From 3921257ecef826177cf55ca1082ff76f100b10ca Mon Sep 17 00:00:00 2001 From: Sebastian Ullrich Date: Thu, 7 Mar 2024 18:02:47 +0100 Subject: [PATCH 01/32] feat: thread initialization for reverse FFI (#3632) Makes it possible to properly allocate and free thread-local runtime resources for threads not started by Lean itself --- doc/dev/ffi.md | 9 +++++++++ src/runtime/thread.cpp | 14 +++++++++++--- src/runtime/thread.h | 5 +++++ 3 files changed, 25 insertions(+), 3 deletions(-) diff --git a/doc/dev/ffi.md b/doc/dev/ffi.md index 98d8e6900586..632965f21682 100644 --- a/doc/dev/ffi.md +++ b/doc/dev/ffi.md @@ -111,6 +111,15 @@ if (lean_io_result_is_ok(res)) { lean_io_mark_end_initialization(); ``` +In addition, any other thread not spawned by the Lean runtime itself must be initialized for Lean use by calling +```c +void lean_initialize_thread(); +``` +and should be finalized in order to free all thread-local resources by calling +```c +void lean_finalize_thread(); +``` + ## `@[extern]` in the Interpreter The interpreter can run Lean declarations for which symbols are available in loaded shared libraries, which includes `@[extern]` declarations. diff --git a/src/runtime/thread.cpp b/src/runtime/thread.cpp index e742fe1e4184..6aab3fbc6d17 100644 --- a/src/runtime/thread.cpp +++ b/src/runtime/thread.cpp @@ -46,18 +46,26 @@ void reset_thread_local() { using runnable = std::function; -static void thread_main(void * p) { +extern "C" LEAN_EXPORT void lean_initialize_thread() { #ifdef LEAN_SMALL_ALLOCATOR init_thread_heap(); #endif +} + +extern "C" LEAN_EXPORT void lean_finalize_thread() { + run_thread_finalizers(); + run_post_thread_finalizers(); +} + +static void thread_main(void * p) { + lean_initialize_thread(); std::unique_ptr f; f.reset(reinterpret_cast(p)); (*f)(); f.reset(); - run_thread_finalizers(); - run_post_thread_finalizers(); + lean_finalize_thread(); } #if defined(LEAN_MULTI_THREAD) diff --git a/src/runtime/thread.h b/src/runtime/thread.h index decebaeae0e8..0d459282025c 100644 --- a/src/runtime/thread.h +++ b/src/runtime/thread.h @@ -217,9 +217,14 @@ static T & GETTER_NAME() { \ } namespace lean { +// module initializer pair (NOT for initializing individual threads!) void initialize_thread(); void finalize_thread(); +// thread initializer pair, for reverse FFI +extern "C" LEAN_EXPORT void lean_initialize_thread(); +extern "C" LEAN_EXPORT void lean_finalize_thread(); + typedef void (*thread_finalizer)(void *); // NOLINT LEAN_EXPORT void register_post_thread_finalizer(thread_finalizer fn, void * p); LEAN_EXPORT void register_thread_finalizer(thread_finalizer fn, void * p); From f336525f3153e278a75876e1f1d3462302c1f23f Mon Sep 17 00:00:00 2001 From: Kyle Miller Date: Thu, 7 Mar 2024 10:14:06 -0800 Subject: [PATCH 02/32] fix: make `delabConstWithSignature` avoid using inaccessible names (#3625) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The `delabConstWithSignature` delaborator is responsible for pretty printing constants with a declaration-like signature, with binders, a colon, and a type. This is used by the `#check` command when it is given just an identifier. It used to accumulate binders from pi types indiscriminately, but this led to unfriendly behavior. For example, `#check String.append` would give ``` String.append (a✝ : String) (a✝¹ : String) : String ``` with inaccessible names. These appear because `String.append` is defined using patterns, so it never names these parameters. Now the delaborator stops accumulating binders once it reaches an inaccessible name, and for example `#check String.append` now gives ``` String.append : String → String → String ``` We do not synthesize names for the sake of enabling binder syntax because the binder names are part of the API of a function — one can use `(arg := ...)` syntax to pass arguments by name. The delaborator also now stops accumulating binders once it reaches a parameter with a name already seen before — we then rely on the main delaborator to provide that parameter with a fresh name when pretty printing the pi type. As a special case, instance parameters with inaccessible names are included as binders, pretty printing like `[LT α]`, rather than relegating them (and all the remaining parameters) to after the colon. It would be more accurate to pretty print this as `[inst✝ : LT α]`, but we make the simplifying assumption that such instance parameters are generally used via typeclass inference. Likely `inst✝` would not directly appear in pretty printer output, and even if it appears in a hover, users can likely figure out what is going on. (We may consider making such `inst✝` variables pretty print as `‹LT α›` or `infer_instance` in the future, to make this more consistent.) Something we note here is that we do not do anything to make sure parameters that can be used as named arguments actually appear named after the colon (nor do we assure that the names are the correct names). For example, one sees `foo : String → String → String` rather than `foo : String → (baz : String) → String`. We can investigate this later if it is wanted. We also give `delabConstWithSignature` a `universes` flag to enable turning off pretty printing universe levels parameters. Closes #2846 --- .../PrettyPrinter/Delaborator/Builtins.lean | 100 +++++++++++++----- tests/lean/1377.lean.expected.out | 2 +- tests/lean/714.lean.expected.out | 2 +- tests/lean/974.lean.expected.out | 16 +-- ...toImplicitChainNameIssue.lean.expected.out | 2 +- .../lean/interactive/hover.lean.expected.out | 8 +- tests/lean/levelNGen.lean.expected.out | 6 +- tests/lean/match1.lean.expected.out | 4 +- tests/lean/namePatEqThm.lean.expected.out | 6 +- tests/lean/run/2846.lean | 74 +++++++++++++ tests/lean/run/funind_demo.lean | 2 +- tests/lean/run/funind_mutual_dep.lean | 2 +- tests/lean/run/funind_tests.lean | 24 +++-- tests/lean/sizeof.lean.expected.out | 3 +- .../lean/terminationFailure.lean.expected.out | 2 +- tests/lean/varBinderUpdate.lean.expected.out | 2 +- 16 files changed, 189 insertions(+), 66 deletions(-) create mode 100644 tests/lean/run/2846.lean diff --git a/src/Lean/PrettyPrinter/Delaborator/Builtins.lean b/src/Lean/PrettyPrinter/Delaborator/Builtins.lean index b34c32b4e4e9..101c17d4f532 100644 --- a/src/Lean/PrettyPrinter/Delaborator/Builtins.lean +++ b/src/Lean/PrettyPrinter/Delaborator/Builtins.lean @@ -1082,41 +1082,85 @@ private unsafe def evalSyntaxConstantUnsafe (env : Environment) (opts : Options) @[implemented_by evalSyntaxConstantUnsafe] private opaque evalSyntaxConstant (env : Environment) (opts : Options) (constName : Name) : ExceptT String Id Syntax := throw "" -/-- Pretty-prints a constant `c` as `c.{} : `. -/ -partial def delabConstWithSignature : Delab := do +/-- +Pretty-prints a constant `c` as `c.{} : `. + +If `universes` is `false`, then the universe level parameters are omitted. +-/ +partial def delabConstWithSignature (universes : Bool := true) : Delab := do let e ← getExpr -- use virtual expression node of arity 2 to separate name and type info let idStx ← descend e 0 <| - withOptions (pp.universes.set · true |> (pp.fullNames.set · true)) <| + withOptions (pp.universes.set · universes |> (pp.fullNames.set · true)) <| delabConst descend (← inferType e) 1 <| - delabParams idStx #[] #[] + delabParams {} idStx #[] where - -- follows `delabBinders`, but does not uniquify binder names and accumulates all binder groups - delabParams (idStx : Ident) (groups : TSyntaxArray ``bracketedBinder) (curIds : Array Ident) := do - if let .forallE n d _ i ← getExpr then - let stxN ← annotateCurPos (mkIdent n) - let curIds := curIds.push ⟨stxN⟩ - if ← shouldGroupWithNext then - withBindingBody n <| delabParams idStx groups curIds - else - let delabTy := withOptions (pp.piBinderTypes.set · true) delab - let group ← withBindingDomain do - match i with - | .implicit => `(bracketedBinderF|{$curIds* : $(← delabTy)}) - | .strictImplicit => `(bracketedBinderF|⦃$curIds* : $(← delabTy)⦄) - | .instImplicit => `(bracketedBinderF|[$curIds.back : $(← delabTy)]) - | _ => - if d.isOptParam then - `(bracketedBinderF|($curIds* : $(← withAppFn <| withAppArg delabTy) := $(← withAppArg delabTy))) - else if let some (.const tacticDecl _) := d.getAutoParamTactic? then - let tacticSyntax ← ofExcept <| evalSyntaxConstant (← getEnv) (← getOptions) tacticDecl - `(bracketedBinderF|($curIds* : $(← withAppFn <| withAppArg delabTy) := by $tacticSyntax)) - else - `(bracketedBinderF|($curIds* : $(← delabTy))) - withBindingBody n <| delabParams idStx (groups.push group) #[] + /-- + For types in the signature, we want to be sure pi binder types are pretty printed. + -/ + delabTy : DelabM Term := withOptions (pp.piBinderTypes.set · true) delab + /- + Similar to `delabBinders`, but does not uniquify binder names (since for named arguments we want to know the name), + and it always merges binder groups when possible. + Once it reaches a binder with an inaccessible name, or a name that has already been used, + the remaining binders appear in pi types after the `:` of the declaration. + -/ + delabParams (bindingNames : NameSet) (idStx : Ident) (groups : TSyntaxArray ``bracketedBinder) := do + let e ← getExpr + if e.isForall && e.binderInfo.isInstImplicit && e.bindingName!.hasMacroScopes then + -- Assumption: this instance can be found by instance search, so it does not need to be named. + -- The oversight here is that this inaccessible name can appear in the pretty printed expression. + -- We could check to see whether the instance appears in the type and avoid omitting the instance name, + -- but this would be the usual case. + let group ← withBindingDomain do `(bracketedBinderF|[$(← delabTy)]) + withBindingBody e.bindingName! <| delabParams bindingNames idStx (groups.push group) + else if e.isForall && !e.bindingName!.hasMacroScopes && !bindingNames.contains e.bindingName! then + delabParamsAux bindingNames idStx groups #[] else - let type ← delab + let type ← delabTy `(declSigWithId| $idStx:ident $groups* : $type) + /-- + Inner loop for `delabParams`, collecting binders. + Invariants: + - The current expression is a forall. + - It has a name that's not inaccessible. + - It has a name that hasn't been used yet. + -/ + delabParamsAux (bindingNames : NameSet) (idStx : Ident) (groups : TSyntaxArray ``bracketedBinder) (curIds : Array Ident) := do + let e@(.forallE n d e' i) ← getExpr | unreachable! + let bindingNames := bindingNames.insert n + let stxN := mkIdent n + let curIds := curIds.push ⟨stxN⟩ + if shouldGroupWithNext bindingNames e e' then + withBindingBody n <| delabParamsAux bindingNames idStx groups curIds + else + let group ← withBindingDomain do + match i with + | .implicit => `(bracketedBinderF|{$curIds* : $(← delabTy)}) + | .strictImplicit => `(bracketedBinderF|⦃$curIds* : $(← delabTy)⦄) + | .instImplicit => `(bracketedBinderF|[$stxN : $(← delabTy)]) + | _ => + if d.isOptParam then + `(bracketedBinderF|($curIds* : $(← withAppFn <| withAppArg delabTy) := $(← withAppArg delabTy))) + else if let some (.const tacticDecl _) := d.getAutoParamTactic? then + let tacticSyntax ← ofExcept <| evalSyntaxConstant (← getEnv) (← getOptions) tacticDecl + `(bracketedBinderF|($curIds* : $(← withAppFn <| withAppArg delabTy) := by $tacticSyntax)) + else + `(bracketedBinderF|($curIds* : $(← delabTy))) + withBindingBody n <| delabParams bindingNames idStx (groups.push group) + /- + Given the forall `e` with body `e'`, determines if the binder from `e'` (if it is a forall) should be grouped with `e`'s binder. + -/ + shouldGroupWithNext (bindingNames : NameSet) (e e' : Expr) : Bool := + e'.isForall && + -- At the first sign of an inaccessible name, stop merging binders: + !e'.bindingName!.hasMacroScopes && + -- If it's a name that has already been used, stop merging binders: + !bindingNames.contains e'.bindingName! && + e.binderInfo == e'.binderInfo && + e.bindingDomain! == e'.bindingDomain! && + -- Inst implicits can't be grouped: + e'.binderInfo != BinderInfo.instImplicit end Lean.PrettyPrinter.Delaborator diff --git a/tests/lean/1377.lean.expected.out b/tests/lean/1377.lean.expected.out index b3ffd2cae82b..b22940f237f4 100644 --- a/tests/lean/1377.lean.expected.out +++ b/tests/lean/1377.lean.expected.out @@ -1,2 +1,2 @@ 1377.lean:3:2-3:5: warning: declaration uses 'sorry' -foo.bar (x✝ : Unit) {n : Nat} (a✝ : Fin n) : Fin n +foo.bar : Unit → {n : Nat} → Fin n → Fin n diff --git a/tests/lean/714.lean.expected.out b/tests/lean/714.lean.expected.out index 3b2a4c436894..de9e1ff1bebe 100644 --- a/tests/lean/714.lean.expected.out +++ b/tests/lean/714.lean.expected.out @@ -1 +1 @@ -test {α : Type} [inst✝ : LT α] : Type +test {α : Type} [LT α] : Type diff --git a/tests/lean/974.lean.expected.out b/tests/lean/974.lean.expected.out index 418dba1e9807..afc0622e0ac4 100644 --- a/tests/lean/974.lean.expected.out +++ b/tests/lean/974.lean.expected.out @@ -1,7 +1,9 @@ -Formula.count_quantifiers._eq_1.{u_1} (x✝ : Nat) (f₁ f₂ : Formula x✝) : - Formula.count_quantifiers (Formula.imp f₁ f₂) = Formula.count_quantifiers f₁ + Formula.count_quantifiers f₂ -Formula.count_quantifiers._eq_2.{u_1} (x✝ : Nat) (f : Formula (x✝ + 1)) : - Formula.count_quantifiers (Formula.all f) = Formula.count_quantifiers f + 1 -Formula.count_quantifiers._eq_3.{u_1} (x✝ : Nat) (x✝¹ : Formula x✝) - (x_4 : ∀ (f₁ f₂ : Formula x✝), x✝¹ = Formula.imp f₁ f₂ → False) - (x_5 : ∀ (f : Formula (x✝ + 1)), x✝¹ = Formula.all f → False) : Formula.count_quantifiers x✝¹ = 0 +Formula.count_quantifiers._eq_1.{u_1} : + ∀ (x : Nat) (f₁ f₂ : Formula x), + Formula.count_quantifiers (Formula.imp f₁ f₂) = Formula.count_quantifiers f₁ + Formula.count_quantifiers f₂ +Formula.count_quantifiers._eq_2.{u_1} : + ∀ (x : Nat) (f : Formula (x + 1)), Formula.count_quantifiers (Formula.all f) = Formula.count_quantifiers f + 1 +Formula.count_quantifiers._eq_3.{u_1} : + ∀ (x : Nat) (x_1 : Formula x), + (∀ (f₁ f₂ : Formula x), x_1 = Formula.imp f₁ f₂ → False) → + (∀ (f : Formula (x + 1)), x_1 = Formula.all f → False) → Formula.count_quantifiers x_1 = 0 diff --git a/tests/lean/autoImplicitChainNameIssue.lean.expected.out b/tests/lean/autoImplicitChainNameIssue.lean.expected.out index e68366ea755e..e6b966055a99 100644 --- a/tests/lean/autoImplicitChainNameIssue.lean.expected.out +++ b/tests/lean/autoImplicitChainNameIssue.lean.expected.out @@ -3,4 +3,4 @@ case nil α✝ : Type u_1 as : List α✝ ⊢ Palindrome (List.reverse []) -palindrome_reverse.{u_1} {α✝ : Type u_1} {as : List α✝} (h : Palindrome as) : Palindrome (List.reverse as) +palindrome_reverse.{u_1} : ∀ {α : Type u_1} {as : List α}, Palindrome as → Palindrome (List.reverse as) diff --git a/tests/lean/interactive/hover.lean.expected.out b/tests/lean/interactive/hover.lean.expected.out index 8b6baeb0700e..5437b2017f25 100644 --- a/tests/lean/interactive/hover.lean.expected.out +++ b/tests/lean/interactive/hover.lean.expected.out @@ -120,7 +120,7 @@ {"start": {"line": 73, "character": 38}, "end": {"line": 73, "character": 45}}, "contents": {"value": - "```lean\nNat.add (a✝a✝¹ : Nat) : Nat\n```\n***\nAddition of natural numbers.\n\nThis definition is overridden in both the kernel and the compiler to efficiently\nevaluate using the \"bignum\" representation (see `Nat`). The definition provided\nhere is the logical model (and it is soundness-critical that they coincide).\n\n***\n*import Init.Prelude*", + "```lean\nNat.add : Nat → Nat → Nat\n```\n***\nAddition of natural numbers.\n\nThis definition is overridden in both the kernel and the compiler to efficiently\nevaluate using the \"bignum\" representation (see `Nat`). The definition provided\nhere is the logical model (and it is soundness-critical that they coincide).\n\n***\n*import Init.Prelude*", "kind": "markdown"}} {"textDocument": {"uri": "file:///hover.lean"}, "position": {"line": 77, "character": 10}} @@ -620,7 +620,7 @@ "end": {"line": 291, "character": 20}}, "contents": {"value": - "```lean\nList.map.{u, v} {α : Type u} {β : Type v} (f : α → β) (a✝ : List α) : List β\n```\n***\n`O(|l|)`. `map f l` applies `f` to each element of the list.\n* `map f [a, b, c] = [f a, f b, f c]`\n\n***\n*import Init.Data.List.Basic*", + "```lean\nList.map.{u, v} {α : Type u} {β : Type v} (f : α → β) : List α → List β\n```\n***\n`O(|l|)`. `map f l` applies `f` to each element of the list.\n* `map f [a, b, c] = [f a, f b, f c]`\n\n***\n*import Init.Data.List.Basic*", "kind": "markdown"}} {"textDocument": {"uri": "file:///hover.lean"}, "position": {"line": 294, "character": 26}} @@ -629,7 +629,7 @@ "end": {"line": 294, "character": 29}}, "contents": {"value": - "```lean\nList.zip.{u, v} {α : Type u} {β : Type v} (a✝ : List α) (a✝¹ : List β) : List (α × β)\n```\n***\n`O(min |xs| |ys|)`. Combines the two lists into a list of pairs, with one element from each list.\nThe longer list is truncated to match the shorter list.\n* `zip [x₁, x₂, x₃] [y₁, y₂, y₃, y₄] = [(x₁, y₁), (x₂, y₂), (x₃, y₃)]`\n\n***\n*import Init.Data.List.Basic*", + "```lean\nList.zip.{u, v} {α : Type u} {β : Type v} : List α → List β → List (α × β)\n```\n***\n`O(min |xs| |ys|)`. Combines the two lists into a list of pairs, with one element from each list.\nThe longer list is truncated to match the shorter list.\n* `zip [x₁, x₂, x₃] [y₁, y₂, y₃, y₄] = [(x₁, y₁), (x₂, y₂), (x₃, y₃)]`\n\n***\n*import Init.Data.List.Basic*", "kind": "markdown"}} {"textDocument": {"uri": "file:///hover.lean"}, "position": {"line": 294, "character": 19}} @@ -638,5 +638,5 @@ "end": {"line": 294, "character": 22}}, "contents": {"value": - "```lean\nList.zip.{u, v} {α : Type u} {β : Type v} (a✝ : List α) (a✝¹ : List β) : List (α × β)\n```\n***\n`O(min |xs| |ys|)`. Combines the two lists into a list of pairs, with one element from each list.\nThe longer list is truncated to match the shorter list.\n* `zip [x₁, x₂, x₃] [y₁, y₂, y₃, y₄] = [(x₁, y₁), (x₂, y₂), (x₃, y₃)]`\n\n***\n*import Init.Data.List.Basic*", + "```lean\nList.zip.{u, v} {α : Type u} {β : Type v} : List α → List β → List (α × β)\n```\n***\n`O(min |xs| |ys|)`. Combines the two lists into a list of pairs, with one element from each list.\nThe longer list is truncated to match the shorter list.\n* `zip [x₁, x₂, x₃] [y₁, y₂, y₃, y₄] = [(x₁, y₁), (x₂, y₂), (x₃, y₃)]`\n\n***\n*import Init.Data.List.Basic*", "kind": "markdown"}} diff --git a/tests/lean/levelNGen.lean.expected.out b/tests/lean/levelNGen.lean.expected.out index 743e0c591930..927f2ce74564 100644 --- a/tests/lean/levelNGen.lean.expected.out +++ b/tests/lean/levelNGen.lean.expected.out @@ -1,5 +1,5 @@ -test1.{u_1, u_2} {α : Sort u_2} (a✝ : α) : Sort u_1 +test1.{u_1, u_2} {α : Sort u_2} : α → Sort u_1 levelNGen.lean:3:4-3:9: warning: declaration uses 'sorry' -test2.{u_1, u_2} {α : Sort u_2} (a✝ : α) : Sort u_1 +test2.{u_1, u_2} {α : Sort u_2} : α → Sort u_1 levelNGen.lean:5:33-5:38: warning: declaration uses 'sorry' -test3.{u_1, u_2} {α : Sort u_2} (a✝ : α) : Sort u_1 +test3.{u_1, u_2} {α : Sort u_2} : α → Sort u_1 diff --git a/tests/lean/match1.lean.expected.out b/tests/lean/match1.lean.expected.out index f7dab2aa43f6..7e2fbddbb467 100644 --- a/tests/lean/match1.lean.expected.out +++ b/tests/lean/match1.lean.expected.out @@ -46,6 +46,6 @@ fun x => | #[] => 0 | #[3, 4, 5] => 3 | x => 4 : Array Nat → Nat -g.match_1.{u_1, u_2} {α : Type u_1} (motive : List α → Sort u_2) (x✝ : List α) (h_1 : (a : α) → motive [a]) - (h_2 : (x : List α) → motive x) : motive x✝ +g.match_1.{u_1, u_2} {α : Type u_1} (motive : List α → Sort u_2) : + (x : List α) → ((a : α) → motive [a]) → ((x : List α) → motive x) → motive x fun e => nomatch e : Empty → False diff --git a/tests/lean/namePatEqThm.lean.expected.out b/tests/lean/namePatEqThm.lean.expected.out index 17bd2acfbad9..73e55b147dc2 100644 --- a/tests/lean/namePatEqThm.lean.expected.out +++ b/tests/lean/namePatEqThm.lean.expected.out @@ -2,5 +2,7 @@ iota._eq_1 : iota 0 = [] iota._eq_2 (n : Nat) : iota (Nat.succ n) = Nat.succ n :: iota n f._eq_1 (x_1 y : Nat) : f [x_1, y] = ([x_1, y], [y]) f._eq_2 (x_1 y : Nat) (zs : List Nat) (x_2 : zs = [] → False) : f (x_1 :: y :: zs) = f zs -f._eq_3 (x✝ : List Nat) (x_2 : ∀ (x y : Nat), x✝ = [x, y] → False) - (x_3 : ∀ (x y : Nat) (zs : List Nat), x✝ = x :: y :: zs → False) : f x✝ = ([], []) +f._eq_3 : + ∀ (x : List Nat), + (∀ (x_1 y : Nat), x = [x_1, y] → False) → + (∀ (x_1 y : Nat) (zs : List Nat), x = x_1 :: y :: zs → False) → f x = ([], []) diff --git a/tests/lean/run/2846.lean b/tests/lean/run/2846.lean new file mode 100644 index 000000000000..d519eb3baea4 --- /dev/null +++ b/tests/lean/run/2846.lean @@ -0,0 +1,74 @@ +/-! +# `delabConstWithSignature` avoids using inaccessible names +-/ + +/-! +Defined without named arguments, prints without named arguments. +-/ +/-- info: String.append : String → String → String -/ +#guard_msgs in #check String.append + +/-! +The List argument is not named, it is not printed as a named argument. +-/ +/-- info: List.length.{u_1} {α : Type u_1} : List α → Nat -/ +#guard_msgs in #check List.length + +/-! +All arguments are named, all are printed as named arguments. +-/ +/-- info: Nat.pos_pow_of_pos {n : Nat} (m : Nat) (h : 0 < n) : 0 < n ^ m -/ +#guard_msgs in #check Nat.pos_pow_of_pos + +/-! +The hypothesis is not a named argument, so it's not printed as a named argument. +-/ +def Nat.pos_pow_of_pos' {n : Nat} (m : Nat) : 0 < n → 0 < n ^ m := Nat.pos_pow_of_pos m + +/-- info: Nat.pos_pow_of_pos' {n : Nat} (m : Nat) : 0 < n → 0 < n ^ m -/ +#guard_msgs in #check Nat.pos_pow_of_pos' + +/-! +Repetition of a named argument, only the first is printed as a named argument. +-/ +def foo (n n : Nat) : Fin (n + 1) := 0 + +/-- info: foo (n : Nat) : (n : Nat) → Fin (n + 1) -/ +#guard_msgs in #check foo + +/-! +Repetition of a named argument, only the first is printed as a named argument. +This is checking that shadowing is handled correctly (that's not the responsibility of +`delabConstWithSignature`, but it assumes that the main delaborator will handle this correctly). +-/ +def foo' (n n : Nat) (a : Fin ((by clear n; exact n) + 1)) : Fin (n + 1) := 0 + +/-- info: foo' (n : Nat) : (n_1 : Nat) → Fin (n + 1) → Fin (n_1 + 1) -/ +#guard_msgs in #check foo' + +/-! +Named argument after inaccesible name, still stays after the colon. +But, it does not print using named pi notation since this is not a dependent type. +-/ +def foo'' : String → (needle : String) → String := fun _ yo => yo + +/-- info: foo'' : String → String → String -/ +#guard_msgs in #check foo'' + +/-! +Named argument after inaccessible name, still stays after the colon. +Here, because it's a dependent type the named pi notation shows the name. +-/ +def bar : String → (n : Nat) → Fin (n+1) := fun _ n => ⟨0, Nat.zero_lt_succ n⟩ + +/-- info: bar : String → (n : Nat) → Fin (n + 1) -/ +#guard_msgs in #check bar + +/-! +Instance argument is an inaccessible name, and we assume that it is a nameless instance, +so it goes before the colon. +-/ +def bar' [LT α] (x : α) : x < x := sorry + +/-- info: bar'.{u_1} {α : Type u_1} [LT α] (x : α) : x < x -/ +#guard_msgs in #check bar' diff --git a/tests/lean/run/funind_demo.lean b/tests/lean/run/funind_demo.lean index e7ea42bfb020..6091baa4f405 100644 --- a/tests/lean/run/funind_demo.lean +++ b/tests/lean/run/funind_demo.lean @@ -10,7 +10,7 @@ derive_functional_induction ackermann info: ackermann.induct (motive : Nat → Nat → Prop) (case1 : ∀ (m : Nat), motive 0 m) (case2 : ∀ (n : Nat), motive n 1 → motive (Nat.succ n) 0) (case3 : ∀ (n m : Nat), motive (n + 1) m → motive n (ackermann (n + 1) m) → motive (Nat.succ n) (Nat.succ m)) - (x x : Nat) : motive x x + (x : Nat) : ∀ (x_1 : Nat), motive x x_1 -/ #guard_msgs in #check ackermann.induct diff --git a/tests/lean/run/funind_mutual_dep.lean b/tests/lean/run/funind_mutual_dep.lean index e3a4d6c96c83..229c26fa423c 100644 --- a/tests/lean/run/funind_mutual_dep.lean +++ b/tests/lean/run/funind_mutual_dep.lean @@ -65,7 +65,7 @@ info: Finite.functions.induct (motive1 : Finite → Prop) (motive2 : (x : Type) (∀ (rest : List (Finite.asType (Finite.arr t1 t2) → α)), motive2 (Finite.asType (Finite.arr t1 t2) → α) t2 rest) → motive2 α (Finite.arr t1 t2) results) - (x : Type) (x : Finite) (x : List x) : motive2 x x x + (x : Type) : ∀ (x_1 : Finite) (x_2 : List x), motive2 x x_1 x_2 -/ #guard_msgs in #check Finite.functions.induct diff --git a/tests/lean/run/funind_tests.lean b/tests/lean/run/funind_tests.lean index 175cf0069d45..1e3cdd526239 100644 --- a/tests/lean/run/funind_tests.lean +++ b/tests/lean/run/funind_tests.lean @@ -32,7 +32,7 @@ derive_functional_induction ackermann info: Binary.ackermann.induct (motive : Nat → Nat → Prop) (case1 : ∀ (m : Nat), motive 0 m) (case2 : ∀ (n : Nat), motive n 1 → motive (Nat.succ n) 0) (case3 : ∀ (n m : Nat), motive (n + 1) m → motive n (ackermann (n + 1) m) → motive (Nat.succ n) (Nat.succ m)) - (x x : Nat) : motive x x + (x : Nat) : ∀ (x_1 : Nat), motive x x_1 -/ #guard_msgs in #check ackermann.induct @@ -244,7 +244,7 @@ derive_functional_induction with_arg_refining_match1 /-- info: with_arg_refining_match1.induct (motive : Nat → Nat → Prop) (case1 : ∀ (i : Nat), motive i 0) (case2 : ∀ (n : Nat), motive 0 (Nat.succ n)) - (case3 : ∀ (i n : Nat), ¬i = 0 → motive (i - 1) n → motive i (Nat.succ n)) (x x : Nat) : motive x x + (case3 : ∀ (i n : Nat), ¬i = 0 → motive (i - 1) n → motive i (Nat.succ n)) (x : Nat) : ∀ (x_1 : Nat), motive x x_1 -/ #guard_msgs in #check with_arg_refining_match1.induct @@ -259,7 +259,8 @@ derive_functional_induction with_arg_refining_match2 /-- info: with_arg_refining_match2.induct (motive : Nat → Nat → Prop) (case1 : ∀ (n : Nat), motive 0 n) (case2 : ∀ (i : Nat), ¬i = 0 → motive i 0) - (case3 : ∀ (i : Nat), ¬i = 0 → ∀ (n : Nat), motive (i - 1) n → motive i (Nat.succ n)) (x x : Nat) : motive x x + (case3 : ∀ (i : Nat), ¬i = 0 → ∀ (n : Nat), motive (i - 1) n → motive i (Nat.succ n)) (x : Nat) : + ∀ (x_1 : Nat), motive x x_1 -/ #guard_msgs in #check with_arg_refining_match2.induct @@ -293,7 +294,8 @@ derive_functional_induction with_mixed_match_tailrec /-- info: with_mixed_match_tailrec.induct (motive : Nat → Nat → Nat → Nat → Prop) (case1 : ∀ (c d x : Nat), motive 0 x c d) - (case2 : ∀ (c d a x : Nat), motive a x (c % 2) (d % 2) → motive (Nat.succ a) x c d) (x x x x : Nat) : motive x x x x + (case2 : ∀ (c d a x : Nat), motive a x (c % 2) (d % 2) → motive (Nat.succ a) x c d) (x : Nat) : + ∀ (x_1 x_2 x_3 : Nat), motive x x_1 x_2 x_3 -/ #guard_msgs in #check with_mixed_match_tailrec.induct @@ -312,8 +314,8 @@ derive_functional_induction with_mixed_match_tailrec2 /-- info: with_mixed_match_tailrec2.induct (motive : Nat → Nat → Nat → Nat → Nat → Prop) (case1 : ∀ (a b c d : Nat), motive 0 a b c d) (case2 : ∀ (c d n x : Nat), motive (Nat.succ n) 0 x c d) - (case3 : ∀ (c d n a x : Nat), motive n a x (c % 2) (d % 2) → motive (Nat.succ n) (Nat.succ a) x c d) - (x x x x x : Nat) : motive x x x x x + (case3 : ∀ (c d n a x : Nat), motive n a x (c % 2) (d % 2) → motive (Nat.succ n) (Nat.succ a) x c d) (x : Nat) : + ∀ (x_1 x_2 x_3 x_4 : Nat), motive x x_1 x_2 x_3 x_4 -/ #guard_msgs in #check with_mixed_match_tailrec2.induct @@ -404,7 +406,7 @@ derive_functional_induction binary /-- info: UnusedExtraParams.binary.induct (base : Nat) (motive : Nat → Nat → Prop) (case1 : ∀ (m : Nat), motive 0 m) - (case2 : ∀ (n m : Nat), motive n m → motive (Nat.succ n) m) (x x : Nat) : motive x x + (case2 : ∀ (n m : Nat), motive n m → motive (Nat.succ n) m) (x : Nat) : ∀ (x_1 : Nat), motive x x_1 -/ #guard_msgs in #check binary.induct @@ -543,7 +545,7 @@ termination_by xs => xs derive_functional_induction foo /-- info: LetFun.foo.induct.{u_1} {α : Type u_1} (x : α) (motive : List α → Prop) (case1 : motive []) - (case2 : ∀ (_y : α) (ys : List α), motive ys → motive (_y :: ys)) (x : List α) : motive x + (case2 : ∀ (_y : α) (ys : List α), motive ys → motive (_y :: ys)) : ∀ (x : List α), motive x -/ #guard_msgs in #check foo.induct @@ -559,7 +561,7 @@ termination_by xs => xs derive_functional_induction bar /-- info: LetFun.bar.induct.{u_1} {α : Type u_1} (x : α) (motive : List α → Prop) (case1 : motive []) - (case2 : ∀ (_y : α) (ys : List α), motive ys → motive (_y :: ys)) (x : List α) : motive x + (case2 : ∀ (_y : α) (ys : List α), motive ys → motive (_y :: ys)) : ∀ (x : List α), motive x -/ #guard_msgs in #check bar.induct @@ -694,7 +696,7 @@ derive_functional_induction foo /-- info: DefaultArgument.foo.induct (fixed : Bool) (motive : Nat → Nat → Prop) (case1 : ∀ (m : Nat), motive 0 m) - (case2 : ∀ (m n : Nat), motive n m → motive (Nat.succ n) m) (x x : Nat) : motive x x + (case2 : ∀ (m n : Nat), motive n m → motive (Nat.succ n) m) (x : Nat) : ∀ (x_1 : Nat), motive x x_1 -/ #guard_msgs in #check foo.induct @@ -721,7 +723,7 @@ info: Nary.foo.induct (motive : Nat → Nat → (x : Nat) → Fin x → Prop) (case5 : ∀ (n m k : Nat) (x : Fin (k + 2)), motive n m (k + 1) { val := 0, isLt := ⋯ } → motive (Nat.succ n) (Nat.succ m) (Nat.succ (Nat.succ k)) x) - (x x x : Nat) (x : Fin x) : motive x x x x + (x : Nat) : ∀ (x_1 x_2 : Nat) (x_3 : Fin x_2), motive x x_1 x_2 x_3 -/ #guard_msgs in #check foo.induct diff --git a/tests/lean/sizeof.lean.expected.out b/tests/lean/sizeof.lean.expected.out index b732b071d9f6..52020debae1f 100644 --- a/tests/lean/sizeof.lean.expected.out +++ b/tests/lean/sizeof.lean.expected.out @@ -7,5 +7,4 @@ 308 310 11 -InfTree.node.sizeOf_spec.{u} {α : Type u} [inst✝ : SizeOf α] (children : Nat → InfTree α) : - sizeOf (InfTree.node children) = 1 +InfTree.node.sizeOf_spec.{u} {α : Type u} [SizeOf α] (children : Nat → InfTree α) : sizeOf (InfTree.node children) = 1 diff --git a/tests/lean/terminationFailure.lean.expected.out b/tests/lean/terminationFailure.lean.expected.out index e2951ec50d76..824aff87fe66 100644 --- a/tests/lean/terminationFailure.lean.expected.out +++ b/tests/lean/terminationFailure.lean.expected.out @@ -16,7 +16,7 @@ x1 = Please use `termination_by` to specify a decreasing measure. f (x : Nat) : Nat -f.g (a✝ : Nat) : Nat +f.g : Nat → Nat 1 2 terminationFailure.lean:24:9-24:12: error: fail to show termination for diff --git a/tests/lean/varBinderUpdate.lean.expected.out b/tests/lean/varBinderUpdate.lean.expected.out index e789dfabe9ec..b23fdb0d8007 100644 --- a/tests/lean/varBinderUpdate.lean.expected.out +++ b/tests/lean/varBinderUpdate.lean.expected.out @@ -1,6 +1,6 @@ f Nat 5 : Nat g 5 : Nat -Ex1.f (α : Type) [inst✝ : Add α] (a : α) : α +Ex1.f (α : Type) [Add α] (a : α) : α Ex1.g {α : Type} (b : α) : α f Nat 5 : Nat g 5 : Nat From ccac989ddab7aa0827c16f355c3cf305aa65fa8a Mon Sep 17 00:00:00 2001 From: Patrick Massot Date: Thu, 7 Mar 2024 15:02:23 -0500 Subject: [PATCH 03/32] doc: expand an error message about compacting closures (#3627) Provide a hint of where the error message may come from. --- src/runtime/compact.cpp | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/runtime/compact.cpp b/src/runtime/compact.cpp index b2269a4cd631..2dcd60d1e355 100644 --- a/src/runtime/compact.cpp +++ b/src/runtime/compact.cpp @@ -342,7 +342,7 @@ void object_compactor::operator()(object * o) { g_tag_counters[lean_ptr_tag(curr)]++; #endif switch (lean_ptr_tag(curr)) { - case LeanClosure: lean_internal_panic("closures cannot be compacted"); + case LeanClosure: lean_internal_panic("closures cannot be compacted. One possible cause of this error is trying to store a function in a persistent environment extension."); case LeanArray: r = insert_array(curr); break; case LeanScalarArray: insert_sarray(curr); break; case LeanString: insert_string(curr); break; From 123dcb964ce8fd5fade53626fd8daa88150fd470 Mon Sep 17 00:00:00 2001 From: Mac Malone Date: Fri, 8 Mar 2024 10:03:07 -0500 Subject: [PATCH 04/32] feat: lake: `LEAN_GITHASH` override (#3609) If the `LEAN_GITHASH` environment variable is set, Lake will now use it instead of the detected Lean's githash when computing traces for builds and the elaborated Lake configuration. This override allows one to replace the Lean version used by a library (e.g., Mathlib) without completely rebuilding it, which is useful for testing custom builds of Lean. --- src/lake/Lake/Build/Monad.lean | 2 +- src/lake/Lake/Config/Env.lean | 17 ++++++++++++++++- src/lake/Lake/Load/Elab.lean | 7 +++---- src/lake/Lake/Load/Main.lean | 10 +++++----- src/lake/tests/env/test.sh | 2 ++ 5 files changed, 27 insertions(+), 11 deletions(-) diff --git a/src/lake/Lake/Build/Monad.lean b/src/lake/Lake/Build/Monad.lean index 9ca09286fc34..cb42d7a082b7 100644 --- a/src/lake/Lake/Build/Monad.lean +++ b/src/lake/Lake/Build/Monad.lean @@ -17,7 +17,7 @@ def mkBuildContext (ws : Workspace) (config : BuildConfig) : IO BuildContext := toBuildConfig := config, startedBuilds := ← IO.mkRef 0 finishedBuilds := ← IO.mkRef 0, - leanTrace := Hash.ofString ws.lakeEnv.lean.githash + leanTrace := Hash.ofString ws.lakeEnv.leanGithash } def failOnBuildCycle [ToString k] : Except (List k) α → BuildM α diff --git a/src/lake/Lake/Config/Env.lean b/src/lake/Lake/Config/Env.lean index 305afc716653..09677b1abdde 100644 --- a/src/lake/Lake/Config/Env.lean +++ b/src/lake/Lake/Config/Env.lean @@ -25,7 +25,9 @@ structure Env where lean : LeanInstall /-- The Elan installation (if any) of the environment. -/ elan? : Option ElanInstall - /-- A name-to-URL mapping of URL overwrites for the named packages. -/ + /-- Overrides the detected Lean's githash as the string Lake uses for Lean traces. -/ + githashOverride : String + /-- A name-to-URL mapping of URL overrides for the named packages. -/ pkgUrlMap : NameMap String /-- The initial Elan toolchain of the environment (i.e., `ELAN_TOOLCHAIN`). -/ initToolchain : String @@ -46,6 +48,7 @@ def compute (lake : LakeInstall) (lean : LeanInstall) (elan? : Option ElanInstal return { lake, lean, elan?, pkgUrlMap := ← computePkgUrlMap + githashOverride := (← IO.getEnv "LEAN_GITHASH").getD "" initToolchain := (← IO.getEnv "ELAN_TOOLCHAIN").getD "" initLeanPath := ← getSearchPath "LEAN_PATH", initLeanSrcPath := ← getSearchPath "LEAN_SRC_PATH", @@ -59,6 +62,17 @@ where | .ok urlMap => return urlMap | .error e => throw s!"'LAKE_PKG_URL_MAP' has invalid JSON: {e}" +/-- +The string Lake uses to identify Lean in traces. +Either the environment-specified `LEAN_GITHASH` or the detected Lean's githash. + +The override allows one to replace the Lean version used by a library +(e.g., Mathlib) without completely rebuilding it, which is useful for testing +custom builds of Lean. +-/ +def leanGithash (env : Env) : String := + if env.githashOverride.isEmpty then env.lean.githash else env.githashOverride + /-- The preferred toolchain of the environment. May be empty. Tries `env.initToolchain` first and then Lake's `Lean.toolchain`. @@ -113,6 +127,7 @@ def baseVars (env : Env) : Array (String × Option String) := ("LAKE", env.lake.lake.toString), ("LAKE_HOME", env.lake.home.toString), ("LAKE_PKG_URL_MAP", toJson env.pkgUrlMap |>.compress), + ("LEAN_GITHASH", env.leanGithash), ("LEAN_SYSROOT", env.lean.sysroot.toString), ("LEAN_AR", env.lean.ar.toString), ("LEAN_CC", env.lean.leanCc?) diff --git a/src/lake/Lake/Load/Elab.lean b/src/lake/Lake/Load/Elab.lean index 9cdf455227f1..c3f64016d12e 100644 --- a/src/lake/Lake/Load/Elab.lean +++ b/src/lake/Lake/Load/Elab.lean @@ -156,7 +156,7 @@ Import the `.olean` for the configuration file if `reconfigure` is not set and an up-to-date one exists (i.e., one with matching configuration and on the same toolchain). Otherwise, elaborate the configuration and save it to the `.olean`. -/ -def importConfigFile (pkgDir lakeDir : FilePath) (lakeOpts : NameMap String) +def importConfigFile (pkgDir lakeDir : FilePath) (lakeEnv : Lake.Env) (lakeOpts : NameMap String) (leanOpts := Options.empty) (configFile := pkgDir / defaultConfigFile) (reconfigure := true) : LogIO Environment := do let some configName := FilePath.mk <$> configFile.fileName | error "invalid configuration file name" @@ -181,7 +181,6 @@ def importConfigFile (pkgDir lakeDir : FilePath) (lakeOpts : NameMap String) handle, and acquires an exclusive lock on the trace. It then releases its lock on the lock file. writes the new lock data. -/ - let acquireTrace h : IO IO.FS.Handle := id do let hLock ← IO.FS.Handle.mk lockFile .write /- @@ -223,7 +222,7 @@ def importConfigFile (pkgDir lakeDir : FilePath) (lakeOpts : NameMap String) match (← IO.FS.removeFile olean |>.toBaseIO) with | .ok _ | .error (.noFileOrDirectory ..) => h.putStrLn <| Json.pretty <| toJson - {platform := System.Platform.target, leanHash := Lean.githash, + {platform := System.Platform.target, leanHash := lakeEnv.leanGithash, configHash, options := lakeOpts : ConfigTrace} h.truncate let env ← elabConfigFile pkgDir lakeOpts leanOpts configFile @@ -245,7 +244,7 @@ def importConfigFile (pkgDir lakeDir : FilePath) (lakeOpts : NameMap String) | error "compiled configuration is invalid; run with '-R' to reconfigure" let upToDate := (← olean.pathExists) ∧ trace.platform = System.Platform.target ∧ - trace.leanHash = Lean.githash ∧ trace.configHash = configHash + trace.leanHash = lakeEnv.leanGithash ∧ trace.configHash = configHash if upToDate then let env ← importConfigFileCore olean leanOpts h.unlock diff --git a/src/lake/Lake/Load/Main.lean b/src/lake/Lake/Load/Main.lean index c865d7c3f2ba..658a99aea69c 100644 --- a/src/lake/Lake/Load/Main.lean +++ b/src/lake/Lake/Load/Main.lean @@ -31,10 +31,10 @@ Elaborate a dependency's configuration file into a `Package`. The resulting package does not yet include any dependencies. -/ def MaterializedDep.loadPackage (dep : MaterializedDep) -(wsDir : FilePath) (leanOpts : Options) (reconfigure : Bool) : LogIO Package := do +(wsDir : FilePath) (lakeEnv : Lake.Env) (leanOpts : Options) (reconfigure : Bool) : LogIO Package := do let dir := wsDir / dep.relPkgDir let lakeDir := dir / defaultLakeDir - let configEnv ← importConfigFile dir lakeDir dep.configOpts leanOpts (dir / dep.configFile) reconfigure + let configEnv ← importConfigFile dir lakeDir lakeEnv dep.configOpts leanOpts (dir / dep.configFile) reconfigure let config ← IO.ofExcept <| PackageConfig.loadFromEnv configEnv leanOpts return { dir @@ -51,7 +51,7 @@ Does not resolve dependencies. def loadWorkspaceRoot (config : LoadConfig) : LogIO Workspace := do Lean.searchPathRef.set config.env.leanSearchPath let configEnv ← importConfigFile - config.rootDir (config.rootDir / defaultLakeDir) + config.rootDir (config.rootDir / defaultLakeDir) config.env config.configOpts config.leanOpts config.configFile config.reconfigure let pkgConfig ← IO.ofExcept <| PackageConfig.loadFromEnv configEnv config.leanOpts let root := { @@ -145,7 +145,7 @@ def Workspace.updateAndMaterialize (ws : Workspace) return pkg else -- Load the package - let depPkg ← dep.loadPackage ws.dir pkg.leanOpts reconfigure + let depPkg ← dep.loadPackage ws.dir ws.lakeEnv pkg.leanOpts reconfigure if depPkg.name ≠ dep.name then logWarning s!"{pkg.name}: package '{depPkg.name}' was required as '{dep.name}'" -- Materialize locked dependencies @@ -220,7 +220,7 @@ def Workspace.materializeDeps (ws : Workspace) (manifest : Manifest) (reconfigur let depPkgs ← deps.mapM fun dep => fetchOrCreate dep.name do if let some entry := pkgEntries.find? dep.name then let result ← entry.materialize dep ws.dir relPkgsDir ws.lakeEnv.pkgUrlMap - result.loadPackage ws.dir pkg.leanOpts reconfigure + result.loadPackage ws.dir ws.lakeEnv pkg.leanOpts reconfigure else if topLevel then error <| s!"dependency '{dep.name}' not in manifest; " ++ diff --git a/src/lake/tests/env/test.sh b/src/lake/tests/env/test.sh index a3036f1d3ba1..4401738c5e5c 100755 --- a/src/lake/tests/env/test.sh +++ b/src/lake/tests/env/test.sh @@ -13,6 +13,7 @@ $LAKE env | grep ".*=.*" # NOTE: `printenv` exits with code 1 if the variable is not set $LAKE env printenv LAKE $LAKE env printenv LAKE_HOME +$LAKE env printenv LEAN_GITHASH $LAKE env printenv LEAN_SYSROOT $LAKE env printenv LEAN_AR | grep ar $LAKE env printenv LEAN_PATH @@ -23,6 +24,7 @@ $LAKE -d ../../examples/hello env printenv PATH | grep examples/hello # Test that `env` preserves the input environment for certain variables test "`$LAKE env env ELAN_TOOLCHAIN=foo $LAKE env printenv ELAN_TOOLCHAIN`" = foo +test "`LEAN_GITHASH=foo $LAKE env printenv LEAN_GITHASH`" = foo test "`LEAN_AR=foo $LAKE env printenv LEAN_AR`" = foo test "`LEAN_CC=foo $LAKE env printenv LEAN_CC`" = foo From 6dd4f4b423a2d8c13fa2a47ffc52c2c94a46638a Mon Sep 17 00:00:00 2001 From: Joe Hendrix Date: Fri, 8 Mar 2024 09:00:46 -0800 Subject: [PATCH 05/32] chore: upstream Std.Data.Nat (#3634) This migrates lemmas about Nat `compare`, `min`, `max`, `dvd`, `gcd`, `lcm` and `div`/`mod` from Std to Lean itself. Std still has some additional recursors, `CoPrime` and a few additional definitions that might merit further discussion prior to upstreaming. --- src/Init/Data/Nat.lean | 2 + src/Init/Data/Nat/Compare.lean | 57 +++++++++++ src/Init/Data/Nat/Div.lean | 46 +++++++++ src/Init/Data/Nat/Dvd.lean | 32 +++++++ src/Init/Data/Nat/Gcd.lean | 170 ++++++++++++++++++++++++++++++++- src/Init/Data/Nat/Lcm.lean | 66 +++++++++++++ src/Init/Data/Nat/Lemmas.lean | 122 +++++++---------------- 7 files changed, 406 insertions(+), 89 deletions(-) create mode 100644 src/Init/Data/Nat/Compare.lean create mode 100644 src/Init/Data/Nat/Lcm.lean diff --git a/src/Init/Data/Nat.lean b/src/Init/Data/Nat.lean index 533194a9e4bd..4d84457ff36e 100644 --- a/src/Init/Data/Nat.lean +++ b/src/Init/Data/Nat.lean @@ -17,3 +17,5 @@ import Init.Data.Nat.Linear import Init.Data.Nat.SOM import Init.Data.Nat.Lemmas import Init.Data.Nat.Mod +import Init.Data.Nat.Lcm +import Init.Data.Nat.Compare diff --git a/src/Init/Data/Nat/Compare.lean b/src/Init/Data/Nat/Compare.lean new file mode 100644 index 000000000000..27914b4233db --- /dev/null +++ b/src/Init/Data/Nat/Compare.lean @@ -0,0 +1,57 @@ +/- +Copyright (c) 2016 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro +-/ +prelude +import Init.Classical +import Init.Data.Ord + +/-! # Basic lemmas about comparing natural numbers + +This file introduce some basic lemmas about compare as applied to natural +numbers. +-/ +namespace Nat + +theorem compare_def_lt (a b : Nat) : + compare a b = if a < b then .lt else if b < a then .gt else .eq := by + simp only [compare, compareOfLessAndEq] + split + · rfl + · next h => + match Nat.lt_or_eq_of_le (Nat.not_lt.1 h) with + | .inl h => simp [h, Nat.ne_of_gt h] + | .inr rfl => simp + +theorem compare_def_le (a b : Nat) : + compare a b = if a ≤ b then if b ≤ a then .eq else .lt else .gt := by + rw [compare_def_lt] + split + · next hlt => simp [Nat.le_of_lt hlt, Nat.not_le.2 hlt] + · next hge => + split + · next hgt => simp [Nat.le_of_lt hgt, Nat.not_le.2 hgt] + · next hle => simp [Nat.not_lt.1 hge, Nat.not_lt.1 hle] + +protected theorem compare_swap (a b : Nat) : (compare a b).swap = compare b a := by + simp only [compare_def_le]; (repeat' split) <;> try rfl + next h1 h2 => cases h1 (Nat.le_of_not_le h2) + +protected theorem compare_eq_eq {a b : Nat} : compare a b = .eq ↔ a = b := by + rw [compare_def_lt]; (repeat' split) <;> simp [Nat.ne_of_lt, Nat.ne_of_gt, *] + next hlt hgt => exact Nat.le_antisymm (Nat.not_lt.1 hgt) (Nat.not_lt.1 hlt) + +protected theorem compare_eq_lt {a b : Nat} : compare a b = .lt ↔ a < b := by + rw [compare_def_lt]; (repeat' split) <;> simp [*] + +protected theorem compare_eq_gt {a b : Nat} : compare a b = .gt ↔ b < a := by + rw [compare_def_lt]; (repeat' split) <;> simp [Nat.le_of_lt, *] + +protected theorem compare_ne_gt {a b : Nat} : compare a b ≠ .gt ↔ a ≤ b := by + rw [compare_def_le]; (repeat' split) <;> simp [*] + +protected theorem compare_ne_lt {a b : Nat} : compare a b ≠ .lt ↔ b ≤ a := by + rw [compare_def_le]; (repeat' split) <;> simp [Nat.le_of_not_le, *] + +end Nat diff --git a/src/Init/Data/Nat/Div.lean b/src/Init/Data/Nat/Div.lean index f0429f4e974d..fbdfb23d1629 100644 --- a/src/Init/Data/Nat/Div.lean +++ b/src/Init/Data/Nat/Div.lean @@ -327,4 +327,50 @@ theorem div_eq_of_lt (h₀ : a < b) : a / b = 0 := by intro h₁ apply Nat.not_le_of_gt h₀ h₁.right +protected theorem mul_div_cancel (m : Nat) {n : Nat} (H : 0 < n) : m * n / n = m := by + let t := add_mul_div_right 0 m H + rwa [Nat.zero_add, Nat.zero_div, Nat.zero_add] at t + +protected theorem mul_div_cancel_left (m : Nat) {n : Nat} (H : 0 < n) : n * m / n = m := by + rw [Nat.mul_comm, Nat.mul_div_cancel _ H] + +protected theorem div_le_of_le_mul {m n : Nat} : ∀ {k}, m ≤ k * n → m / k ≤ n + | 0, _ => by simp [Nat.div_zero, n.zero_le] + | succ k, h => by + suffices succ k * (m / succ k) ≤ succ k * n from + Nat.le_of_mul_le_mul_left this (zero_lt_succ _) + have h1 : succ k * (m / succ k) ≤ m % succ k + succ k * (m / succ k) := Nat.le_add_left _ _ + have h2 : m % succ k + succ k * (m / succ k) = m := by rw [mod_add_div] + have h3 : m ≤ succ k * n := h + rw [← h2] at h3 + exact Nat.le_trans h1 h3 + +@[simp] theorem mul_div_right (n : Nat) {m : Nat} (H : 0 < m) : m * n / m = n := by + induction n <;> simp_all [mul_succ] + +@[simp] theorem mul_div_left (m : Nat) {n : Nat} (H : 0 < n) : m * n / n = m := by + rw [Nat.mul_comm, mul_div_right _ H] + +protected theorem div_self (H : 0 < n) : n / n = 1 := by + let t := add_div_right 0 H + rwa [Nat.zero_add, Nat.zero_div] at t + +protected theorem div_eq_of_eq_mul_left (H1 : 0 < n) (H2 : m = k * n) : m / n = k := +by rw [H2, Nat.mul_div_cancel _ H1] + +protected theorem div_eq_of_eq_mul_right (H1 : 0 < n) (H2 : m = n * k) : m / n = k := +by rw [H2, Nat.mul_div_cancel_left _ H1] + +protected theorem mul_div_mul_left {m : Nat} (n k : Nat) (H : 0 < m) : + m * n / (m * k) = n / k := by rw [← Nat.div_div_eq_div_mul, Nat.mul_div_cancel_left _ H] + +protected theorem mul_div_mul_right {m : Nat} (n k : Nat) (H : 0 < m) : + n * m / (k * m) = n / k := by rw [Nat.mul_comm, Nat.mul_comm k, Nat.mul_div_mul_left _ _ H] + +theorem mul_div_le (m n : Nat) : n * (m / n) ≤ m := by + match n, Nat.eq_zero_or_pos n with + | _, Or.inl rfl => rw [Nat.zero_mul]; exact m.zero_le + | n, Or.inr h => rw [Nat.mul_comm, ← Nat.le_div_iff_mul_le h]; exact Nat.le_refl _ + + end Nat diff --git a/src/Init/Data/Nat/Dvd.lean b/src/Init/Data/Nat/Dvd.lean index 1b35d41d5fb9..ad3eff53089e 100644 --- a/src/Init/Data/Nat/Dvd.lean +++ b/src/Init/Data/Nat/Dvd.lean @@ -104,4 +104,36 @@ protected theorem div_mul_cancel {n m : Nat} (H : n ∣ m) : m / n * n = m := by subst h rw [Nat.mul_assoc, add_mul_mod_self_left] +protected theorem dvd_of_mul_dvd_mul_left + (kpos : 0 < k) (H : k * m ∣ k * n) : m ∣ n := by + let ⟨l, H⟩ := H + rw [Nat.mul_assoc] at H + exact ⟨_, Nat.eq_of_mul_eq_mul_left kpos H⟩ + +protected theorem dvd_of_mul_dvd_mul_right (kpos : 0 < k) (H : m * k ∣ n * k) : m ∣ n := by + rw [Nat.mul_comm m k, Nat.mul_comm n k] at H; exact Nat.dvd_of_mul_dvd_mul_left kpos H + +theorem dvd_sub {k m n : Nat} (H : n ≤ m) (h₁ : k ∣ m) (h₂ : k ∣ n) : k ∣ m - n := + (Nat.dvd_add_iff_left h₂).2 <| by rwa [Nat.sub_add_cancel H] + +protected theorem mul_dvd_mul {a b c d : Nat} : a ∣ b → c ∣ d → a * c ∣ b * d + | ⟨e, he⟩, ⟨f, hf⟩ => + ⟨e * f, by simp [he, hf, Nat.mul_assoc, Nat.mul_left_comm, Nat.mul_comm]⟩ + +protected theorem mul_dvd_mul_left (a : Nat) (h : b ∣ c) : a * b ∣ a * c := + Nat.mul_dvd_mul (Nat.dvd_refl a) h + +protected theorem mul_dvd_mul_right (h: a ∣ b) (c : Nat) : a * c ∣ b * c := + Nat.mul_dvd_mul h (Nat.dvd_refl c) + +@[simp] theorem dvd_one {n : Nat} : n ∣ 1 ↔ n = 1 := + ⟨eq_one_of_dvd_one, fun h => h.symm ▸ Nat.dvd_refl _⟩ + +protected theorem mul_div_assoc (m : Nat) (H : k ∣ n) : m * n / k = m * (n / k) := by + match Nat.eq_zero_or_pos k with + | .inl h0 => rw [h0, Nat.div_zero, Nat.div_zero, Nat.mul_zero] + | .inr hpos => + have h1 : m * n / k = m * (n / k * k) / k := by rw [Nat.div_mul_cancel H] + rw [h1, ← Nat.mul_assoc, Nat.mul_div_cancel _ hpos] + end Nat diff --git a/src/Init/Data/Nat/Gcd.lean b/src/Init/Data/Nat/Gcd.lean index cb66f767fb02..f03cf327b1b4 100644 --- a/src/Init/Data/Nat/Gcd.lean +++ b/src/Init/Data/Nat/Gcd.lean @@ -1,10 +1,12 @@ /- Copyright (c) 2021 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Leonardo de Moura +Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro -/ prelude import Init.Data.Nat.Dvd +import Init.NotationExtra +import Init.RCases namespace Nat @@ -14,8 +16,8 @@ def gcd (m n : @& Nat) : Nat := n else gcd (n % m) m -termination_by m -decreasing_by simp_wf; apply mod_lt _ (zero_lt_of_ne_zero _); assumption + termination_by m + decreasing_by simp_wf; apply mod_lt _ (zero_lt_of_ne_zero _); assumption @[simp] theorem gcd_zero_left (y : Nat) : gcd 0 y = y := rfl @@ -69,4 +71,166 @@ theorem dvd_gcd : k ∣ m → k ∣ n → k ∣ gcd m n := by | H0 n => rw [gcd_zero_left]; exact kn | H1 n m _ IH => rw [gcd_rec]; exact IH ((dvd_mod_iff km).2 kn) km +theorem dvd_gcd_iff : k ∣ gcd m n ↔ k ∣ m ∧ k ∣ n := + ⟨fun h => let ⟨h₁, h₂⟩ := gcd_dvd m n; ⟨Nat.dvd_trans h h₁, Nat.dvd_trans h h₂⟩, + fun ⟨h₁, h₂⟩ => dvd_gcd h₁ h₂⟩ + +theorem gcd_comm (m n : Nat) : gcd m n = gcd n m := + Nat.dvd_antisymm + (dvd_gcd (gcd_dvd_right m n) (gcd_dvd_left m n)) + (dvd_gcd (gcd_dvd_right n m) (gcd_dvd_left n m)) + +theorem gcd_eq_left_iff_dvd : m ∣ n ↔ gcd m n = m := + ⟨fun h => by rw [gcd_rec, mod_eq_zero_of_dvd h, gcd_zero_left], + fun h => h ▸ gcd_dvd_right m n⟩ + +theorem gcd_eq_right_iff_dvd : m ∣ n ↔ gcd n m = m := by + rw [gcd_comm]; exact gcd_eq_left_iff_dvd + +theorem gcd_assoc (m n k : Nat) : gcd (gcd m n) k = gcd m (gcd n k) := + Nat.dvd_antisymm + (dvd_gcd + (Nat.dvd_trans (gcd_dvd_left (gcd m n) k) (gcd_dvd_left m n)) + (dvd_gcd (Nat.dvd_trans (gcd_dvd_left (gcd m n) k) (gcd_dvd_right m n)) + (gcd_dvd_right (gcd m n) k))) + (dvd_gcd + (dvd_gcd (gcd_dvd_left m (gcd n k)) + (Nat.dvd_trans (gcd_dvd_right m (gcd n k)) (gcd_dvd_left n k))) + (Nat.dvd_trans (gcd_dvd_right m (gcd n k)) (gcd_dvd_right n k))) + +@[simp] theorem gcd_one_right (n : Nat) : gcd n 1 = 1 := (gcd_comm n 1).trans (gcd_one_left n) + +theorem gcd_mul_left (m n k : Nat) : gcd (m * n) (m * k) = m * gcd n k := by + induction n, k using gcd.induction with + | H0 k => simp + | H1 n k _ IH => rwa [← mul_mod_mul_left, ← gcd_rec, ← gcd_rec] at IH + +theorem gcd_mul_right (m n k : Nat) : gcd (m * n) (k * n) = gcd m k * n := by + rw [Nat.mul_comm m n, Nat.mul_comm k n, Nat.mul_comm (gcd m k) n, gcd_mul_left] + +theorem gcd_pos_of_pos_left {m : Nat} (n : Nat) (mpos : 0 < m) : 0 < gcd m n := + pos_of_dvd_of_pos (gcd_dvd_left m n) mpos + +theorem gcd_pos_of_pos_right (m : Nat) {n : Nat} (npos : 0 < n) : 0 < gcd m n := + pos_of_dvd_of_pos (gcd_dvd_right m n) npos + +theorem div_gcd_pos_of_pos_left (b : Nat) (h : 0 < a) : 0 < a / a.gcd b := + (Nat.le_div_iff_mul_le <| Nat.gcd_pos_of_pos_left _ h).2 (Nat.one_mul _ ▸ Nat.gcd_le_left _ h) + +theorem div_gcd_pos_of_pos_right (a : Nat) (h : 0 < b) : 0 < b / a.gcd b := + (Nat.le_div_iff_mul_le <| Nat.gcd_pos_of_pos_right _ h).2 (Nat.one_mul _ ▸ Nat.gcd_le_right _ h) + +theorem eq_zero_of_gcd_eq_zero_left {m n : Nat} (H : gcd m n = 0) : m = 0 := + match eq_zero_or_pos m with + | .inl H0 => H0 + | .inr H1 => absurd (Eq.symm H) (ne_of_lt (gcd_pos_of_pos_left _ H1)) + +theorem eq_zero_of_gcd_eq_zero_right {m n : Nat} (H : gcd m n = 0) : n = 0 := by + rw [gcd_comm] at H + exact eq_zero_of_gcd_eq_zero_left H + +theorem gcd_ne_zero_left : m ≠ 0 → gcd m n ≠ 0 := mt eq_zero_of_gcd_eq_zero_left + +theorem gcd_ne_zero_right : n ≠ 0 → gcd m n ≠ 0 := mt eq_zero_of_gcd_eq_zero_right + +theorem gcd_div {m n k : Nat} (H1 : k ∣ m) (H2 : k ∣ n) : + gcd (m / k) (n / k) = gcd m n / k := + match eq_zero_or_pos k with + | .inl H0 => by simp [H0] + | .inr H3 => by + apply Nat.eq_of_mul_eq_mul_right H3 + rw [Nat.div_mul_cancel (dvd_gcd H1 H2), ← gcd_mul_right, + Nat.div_mul_cancel H1, Nat.div_mul_cancel H2] + +theorem gcd_dvd_gcd_of_dvd_left {m k : Nat} (n : Nat) (H : m ∣ k) : gcd m n ∣ gcd k n := + dvd_gcd (Nat.dvd_trans (gcd_dvd_left m n) H) (gcd_dvd_right m n) + +theorem gcd_dvd_gcd_of_dvd_right {m k : Nat} (n : Nat) (H : m ∣ k) : gcd n m ∣ gcd n k := + dvd_gcd (gcd_dvd_left n m) (Nat.dvd_trans (gcd_dvd_right n m) H) + +theorem gcd_dvd_gcd_mul_left (m n k : Nat) : gcd m n ∣ gcd (k * m) n := + gcd_dvd_gcd_of_dvd_left _ (Nat.dvd_mul_left _ _) + +theorem gcd_dvd_gcd_mul_right (m n k : Nat) : gcd m n ∣ gcd (m * k) n := + gcd_dvd_gcd_of_dvd_left _ (Nat.dvd_mul_right _ _) + +theorem gcd_dvd_gcd_mul_left_right (m n k : Nat) : gcd m n ∣ gcd m (k * n) := + gcd_dvd_gcd_of_dvd_right _ (Nat.dvd_mul_left _ _) + +theorem gcd_dvd_gcd_mul_right_right (m n k : Nat) : gcd m n ∣ gcd m (n * k) := + gcd_dvd_gcd_of_dvd_right _ (Nat.dvd_mul_right _ _) + +theorem gcd_eq_left {m n : Nat} (H : m ∣ n) : gcd m n = m := + Nat.dvd_antisymm (gcd_dvd_left _ _) (dvd_gcd (Nat.dvd_refl _) H) + +theorem gcd_eq_right {m n : Nat} (H : n ∣ m) : gcd m n = n := by + rw [gcd_comm, gcd_eq_left H] + +@[simp] theorem gcd_mul_left_left (m n : Nat) : gcd (m * n) n = n := + Nat.dvd_antisymm (gcd_dvd_right _ _) (dvd_gcd (Nat.dvd_mul_left _ _) (Nat.dvd_refl _)) + +@[simp] theorem gcd_mul_left_right (m n : Nat) : gcd n (m * n) = n := by + rw [gcd_comm, gcd_mul_left_left] + +@[simp] theorem gcd_mul_right_left (m n : Nat) : gcd (n * m) n = n := by + rw [Nat.mul_comm, gcd_mul_left_left] + +@[simp] theorem gcd_mul_right_right (m n : Nat) : gcd n (n * m) = n := by + rw [gcd_comm, gcd_mul_right_left] + +@[simp] theorem gcd_gcd_self_right_left (m n : Nat) : gcd m (gcd m n) = gcd m n := + Nat.dvd_antisymm (gcd_dvd_right _ _) (dvd_gcd (gcd_dvd_left _ _) (Nat.dvd_refl _)) + +@[simp] theorem gcd_gcd_self_right_right (m n : Nat) : gcd m (gcd n m) = gcd n m := by + rw [gcd_comm n m, gcd_gcd_self_right_left] + +@[simp] theorem gcd_gcd_self_left_right (m n : Nat) : gcd (gcd n m) m = gcd n m := by + rw [gcd_comm, gcd_gcd_self_right_right] + +@[simp] theorem gcd_gcd_self_left_left (m n : Nat) : gcd (gcd m n) m = gcd m n := by + rw [gcd_comm m n, gcd_gcd_self_left_right] + +theorem gcd_add_mul_self (m n k : Nat) : gcd m (n + k * m) = gcd m n := by + simp [gcd_rec m (n + k * m), gcd_rec m n] + +theorem gcd_eq_zero_iff {i j : Nat} : gcd i j = 0 ↔ i = 0 ∧ j = 0 := + ⟨fun h => ⟨eq_zero_of_gcd_eq_zero_left h, eq_zero_of_gcd_eq_zero_right h⟩, + fun h => by simp [h]⟩ + +/-- Characterization of the value of `Nat.gcd`. -/ +theorem gcd_eq_iff (a b : Nat) : + gcd a b = g ↔ g ∣ a ∧ g ∣ b ∧ (∀ c, c ∣ a → c ∣ b → c ∣ g) := by + constructor + · rintro rfl + exact ⟨gcd_dvd_left _ _, gcd_dvd_right _ _, fun _ => Nat.dvd_gcd⟩ + · rintro ⟨ha, hb, hc⟩ + apply Nat.dvd_antisymm + · apply hc + · exact gcd_dvd_left a b + · exact gcd_dvd_right a b + · exact Nat.dvd_gcd ha hb + +/-- Represent a divisor of `m * n` as a product of a divisor of `m` and a divisor of `n`. -/ +def prod_dvd_and_dvd_of_dvd_prod {k m n : Nat} (H : k ∣ m * n) : + {d : {m' // m' ∣ m} × {n' // n' ∣ n} // k = d.1.val * d.2.val} := + if h0 : gcd k m = 0 then + ⟨⟨⟨0, eq_zero_of_gcd_eq_zero_right h0 ▸ Nat.dvd_refl 0⟩, + ⟨n, Nat.dvd_refl n⟩⟩, + eq_zero_of_gcd_eq_zero_left h0 ▸ (Nat.zero_mul n).symm⟩ + else by + have hd : gcd k m * (k / gcd k m) = k := Nat.mul_div_cancel' (gcd_dvd_left k m) + refine ⟨⟨⟨gcd k m, gcd_dvd_right k m⟩, ⟨k / gcd k m, ?_⟩⟩, hd.symm⟩ + apply Nat.dvd_of_mul_dvd_mul_left (Nat.pos_of_ne_zero h0) + rw [hd, ← gcd_mul_right] + exact Nat.dvd_gcd (Nat.dvd_mul_right _ _) H + +theorem gcd_mul_dvd_mul_gcd (k m n : Nat) : gcd k (m * n) ∣ gcd k m * gcd k n := by + let ⟨⟨⟨m', hm'⟩, ⟨n', hn'⟩⟩, (h : gcd k (m * n) = m' * n')⟩ := + prod_dvd_and_dvd_of_dvd_prod <| gcd_dvd_right k (m * n) + rw [h] + have h' : m' * n' ∣ k := h ▸ gcd_dvd_left .. + exact Nat.mul_dvd_mul + (dvd_gcd (Nat.dvd_trans (Nat.dvd_mul_right m' n') h') hm') + (dvd_gcd (Nat.dvd_trans (Nat.dvd_mul_left n' m') h') hn') + end Nat diff --git a/src/Init/Data/Nat/Lcm.lean b/src/Init/Data/Nat/Lcm.lean new file mode 100644 index 000000000000..4573c26d6d0e --- /dev/null +++ b/src/Init/Data/Nat/Lcm.lean @@ -0,0 +1,66 @@ +/- +Copyright (c) 2014 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro +-/ +prelude +import Init.Data.Nat.Gcd +import Init.Data.Nat.Lemmas + +namespace Nat + +/-- The least common multiple of `m` and `n`, defined using `gcd`. -/ +def lcm (m n : Nat) : Nat := m * n / gcd m n + +theorem lcm_comm (m n : Nat) : lcm m n = lcm n m := by + rw [lcm, lcm, Nat.mul_comm n m, gcd_comm n m] + +@[simp] theorem lcm_zero_left (m : Nat) : lcm 0 m = 0 := by simp [lcm] + +@[simp] theorem lcm_zero_right (m : Nat) : lcm m 0 = 0 := by simp [lcm] + +@[simp] theorem lcm_one_left (m : Nat) : lcm 1 m = m := by simp [lcm] + +@[simp] theorem lcm_one_right (m : Nat) : lcm m 1 = m := by simp [lcm] + +@[simp] theorem lcm_self (m : Nat) : lcm m m = m := by + match eq_zero_or_pos m with + | .inl h => rw [h, lcm_zero_left] + | .inr h => simp [lcm, Nat.mul_div_cancel _ h] + +theorem dvd_lcm_left (m n : Nat) : m ∣ lcm m n := + ⟨n / gcd m n, by rw [← Nat.mul_div_assoc m (Nat.gcd_dvd_right m n)]; rfl⟩ + +theorem dvd_lcm_right (m n : Nat) : n ∣ lcm m n := lcm_comm n m ▸ dvd_lcm_left n m + +theorem gcd_mul_lcm (m n : Nat) : gcd m n * lcm m n = m * n := by + rw [lcm, Nat.mul_div_cancel' (Nat.dvd_trans (gcd_dvd_left m n) (Nat.dvd_mul_right m n))] + +theorem lcm_dvd {m n k : Nat} (H1 : m ∣ k) (H2 : n ∣ k) : lcm m n ∣ k := by + match eq_zero_or_pos k with + | .inl h => rw [h]; exact Nat.dvd_zero _ + | .inr kpos => + apply Nat.dvd_of_mul_dvd_mul_left (gcd_pos_of_pos_left n (pos_of_dvd_of_pos H1 kpos)) + rw [gcd_mul_lcm, ← gcd_mul_right, Nat.mul_comm n k] + exact dvd_gcd (Nat.mul_dvd_mul_left _ H2) (Nat.mul_dvd_mul_right H1 _) + +theorem lcm_assoc (m n k : Nat) : lcm (lcm m n) k = lcm m (lcm n k) := +Nat.dvd_antisymm + (lcm_dvd + (lcm_dvd (dvd_lcm_left m (lcm n k)) + (Nat.dvd_trans (dvd_lcm_left n k) (dvd_lcm_right m (lcm n k)))) + (Nat.dvd_trans (dvd_lcm_right n k) (dvd_lcm_right m (lcm n k)))) + (lcm_dvd + (Nat.dvd_trans (dvd_lcm_left m n) (dvd_lcm_left (lcm m n) k)) + (lcm_dvd (Nat.dvd_trans (dvd_lcm_right m n) (dvd_lcm_left (lcm m n) k)) + (dvd_lcm_right (lcm m n) k))) + +theorem lcm_ne_zero (hm : m ≠ 0) (hn : n ≠ 0) : lcm m n ≠ 0 := by + intro h + have h1 := gcd_mul_lcm m n + rw [h, Nat.mul_zero] at h1 + match mul_eq_zero.1 h1.symm with + | .inl hm1 => exact hm hm1 + | .inr hn1 => exact hn hn1 + +end Nat diff --git a/src/Init/Data/Nat/Lemmas.lean b/src/Init/Data/Nat/Lemmas.lean index c35c5c547873..8ffa17d876d2 100644 --- a/src/Init/Data/Nat/Lemmas.lean +++ b/src/Init/Data/Nat/Lemmas.lean @@ -8,6 +8,7 @@ import Init.Data.Nat.Dvd import Init.Data.Nat.MinMax import Init.Data.Nat.Log2 import Init.Data.Nat.Power2 +import Init.Omega /-! # Basic lemmas about natural numbers @@ -335,6 +336,32 @@ protected theorem sub_max_sub_right : ∀ (a b c : Nat), max (a - c) (b - c) = m | _, _, 0 => rfl | _, _, _+1 => Eq.trans (Nat.pred_max_pred ..) <| congrArg _ (Nat.sub_max_sub_right ..) +protected theorem sub_min_sub_left (a b c : Nat) : min (a - b) (a - c) = a - max b c := by + omega + +protected theorem sub_max_sub_left (a b c : Nat) : max (a - b) (a - c) = a - min b c := by + omega + +protected theorem mul_max_mul_right (a b c : Nat) : max (a * c) (b * c) = max a b * c := by + induction a generalizing b with + | zero => simp + | succ i ind => + cases b <;> simp [succ_eq_add_one, Nat.succ_mul, Nat.add_max_add_right, ind] + +protected theorem mul_min_mul_right (a b c : Nat) : min (a * c) (b * c) = min a b * c := by + induction a generalizing b with + | zero => simp + | succ i ind => + cases b <;> simp [succ_eq_add_one, Nat.succ_mul, Nat.add_min_add_right, ind] + +protected theorem mul_max_mul_left (a b c : Nat) : max (a * b) (a * c) = a * max b c := by + repeat rw [Nat.mul_comm a] + exact Nat.mul_max_mul_right .. + +protected theorem mul_min_mul_left (a b c : Nat) : min (a * b) (a * c) = a * min b c := by + repeat rw [Nat.mul_comm a] + exact Nat.mul_min_mul_right .. + -- protected theorem sub_min_sub_left (a b c : Nat) : min (a - b) (a - c) = a - max b c := by -- induction b, c using Nat.recDiagAux with -- | zero_left => rw [Nat.sub_zero, Nat.zero_max]; exact Nat.min_eq_right (Nat.sub_le ..) @@ -484,51 +511,6 @@ protected theorem pos_of_mul_pos_right {a b : Nat} (h : 0 < a * b) : 0 < a := by /-! ### div/mod -/ -protected theorem div_le_of_le_mul {m n : Nat} : ∀ {k}, m ≤ k * n → m / k ≤ n - | 0, _ => by simp [Nat.div_zero, n.zero_le] - | succ k, h => by - suffices succ k * (m / succ k) ≤ succ k * n from - Nat.le_of_mul_le_mul_left this (zero_lt_succ _) - have h1 : succ k * (m / succ k) ≤ m % succ k + succ k * (m / succ k) := Nat.le_add_left _ _ - have h2 : m % succ k + succ k * (m / succ k) = m := by rw [mod_add_div] - have h3 : m ≤ succ k * n := h - rw [← h2] at h3 - exact Nat.le_trans h1 h3 - -@[simp] theorem mul_div_right (n : Nat) {m : Nat} (H : 0 < m) : m * n / m = n := by - induction n <;> simp_all [mul_succ] - -@[simp] theorem mul_div_left (m : Nat) {n : Nat} (H : 0 < n) : m * n / n = m := by - rw [Nat.mul_comm, mul_div_right _ H] - -protected theorem div_self (H : 0 < n) : n / n = 1 := by - let t := add_div_right 0 H - rwa [Nat.zero_add, Nat.zero_div] at t - -protected theorem mul_div_cancel (m : Nat) {n : Nat} (H : 0 < n) : m * n / n = m := by - let t := add_mul_div_right 0 m H - rwa [Nat.zero_add, Nat.zero_div, Nat.zero_add] at t - -protected theorem mul_div_cancel_left (m : Nat) {n : Nat} (H : 0 < n) : n * m / n = m := -by rw [Nat.mul_comm, Nat.mul_div_cancel _ H] - -protected theorem div_eq_of_eq_mul_left (H1 : 0 < n) (H2 : m = k * n) : m / n = k := -by rw [H2, Nat.mul_div_cancel _ H1] - -protected theorem div_eq_of_eq_mul_right (H1 : 0 < n) (H2 : m = n * k) : m / n = k := -by rw [H2, Nat.mul_div_cancel_left _ H1] - -protected theorem mul_div_mul_left {m : Nat} (n k : Nat) (H : 0 < m) : - m * n / (m * k) = n / k := by rw [← Nat.div_div_eq_div_mul, Nat.mul_div_cancel_left _ H] - -protected theorem mul_div_mul_right {m : Nat} (n k : Nat) (H : 0 < m) : - n * m / (k * m) = n / k := by rw [Nat.mul_comm, Nat.mul_comm k, Nat.mul_div_mul_left _ _ H] - -theorem mul_div_le (m n : Nat) : n * (m / n) ≤ m := by - match n, Nat.eq_zero_or_pos n with - | _, Or.inl rfl => rw [Nat.zero_mul]; exact m.zero_le - | n, Or.inr h => rw [Nat.mul_comm, ← Nat.le_div_iff_mul_le h]; exact Nat.le_refl _ - theorem mod_two_eq_zero_or_one (n : Nat) : n % 2 = 0 ∨ n % 2 = 1 := match n % 2, @Nat.mod_lt n 2 (by decide) with | 0, _ => .inl rfl @@ -719,37 +701,17 @@ theorem lt_log2_self : n < 2 ^ (n.log2 + 1) := /-! ### dvd -/ -theorem dvd_sub {k m n : Nat} (H : n ≤ m) (h₁ : k ∣ m) (h₂ : k ∣ n) : k ∣ m - n := - (Nat.dvd_add_iff_left h₂).2 <| by rwa [Nat.sub_add_cancel H] - -protected theorem mul_dvd_mul {a b c d : Nat} : a ∣ b → c ∣ d → a * c ∣ b * d - | ⟨e, he⟩, ⟨f, hf⟩ => - ⟨e * f, by simp [he, hf, Nat.mul_assoc, Nat.mul_left_comm, Nat.mul_comm]⟩ - -protected theorem mul_dvd_mul_left (a : Nat) (h : b ∣ c) : a * b ∣ a * c := - Nat.mul_dvd_mul (Nat.dvd_refl a) h - -protected theorem mul_dvd_mul_right (h: a ∣ b) (c : Nat) : a * c ∣ b * c := - Nat.mul_dvd_mul h (Nat.dvd_refl c) - -@[simp] theorem dvd_one {n : Nat} : n ∣ 1 ↔ n = 1 := - ⟨eq_one_of_dvd_one, fun h => h.symm ▸ Nat.dvd_refl _⟩ - -protected theorem mul_div_assoc (m : Nat) (H : k ∣ n) : m * n / k = m * (n / k) := by - match Nat.eq_zero_or_pos k with - | .inl h0 => rw [h0, Nat.div_zero, Nat.div_zero, Nat.mul_zero] - | .inr hpos => - have h1 : m * n / k = m * (n / k * k) / k := by rw [Nat.div_mul_cancel H] - rw [h1, ← Nat.mul_assoc, Nat.mul_div_cancel _ hpos] +protected theorem eq_mul_of_div_eq_right {a b c : Nat} (H1 : b ∣ a) (H2 : a / b = c) : + a = b * c := by + rw [← H2, Nat.mul_div_cancel' H1] -protected theorem dvd_of_mul_dvd_mul_left - (kpos : 0 < k) (H : k * m ∣ k * n) : m ∣ n := by - let ⟨l, H⟩ := H - rw [Nat.mul_assoc] at H - exact ⟨_, Nat.eq_of_mul_eq_mul_left kpos H⟩ +protected theorem div_eq_iff_eq_mul_right {a b c : Nat} (H : 0 < b) (H' : b ∣ a) : + a / b = c ↔ a = b * c := + ⟨Nat.eq_mul_of_div_eq_right H', Nat.div_eq_of_eq_mul_right H⟩ -protected theorem dvd_of_mul_dvd_mul_right (kpos : 0 < k) (H : m * k ∣ n * k) : m ∣ n := by - rw [Nat.mul_comm m k, Nat.mul_comm n k] at H; exact Nat.dvd_of_mul_dvd_mul_left kpos H +protected theorem div_eq_iff_eq_mul_left {a b c : Nat} (H : 0 < b) (H' : b ∣ a) : + a / b = c ↔ a = c * b := by + rw [Nat.mul_comm]; exact Nat.div_eq_iff_eq_mul_right H H' theorem pow_dvd_pow_iff_pow_le_pow {k l : Nat} : ∀ {x : Nat}, 0 < x → (x ^ k ∣ x ^ l ↔ x ^ k ≤ x ^ l) @@ -773,18 +735,6 @@ theorem pow_dvd_pow_iff_le_right {x k l : Nat} (w : 1 < x) : x ^ k ∣ x ^ l ↔ theorem pow_dvd_pow_iff_le_right' {b k l : Nat} : (b + 2) ^ k ∣ (b + 2) ^ l ↔ k ≤ l := pow_dvd_pow_iff_le_right (Nat.lt_of_sub_eq_succ rfl) -protected theorem eq_mul_of_div_eq_right {a b c : Nat} (H1 : b ∣ a) (H2 : a / b = c) : - a = b * c := by - rw [← H2, Nat.mul_div_cancel' H1] - -protected theorem div_eq_iff_eq_mul_right {a b c : Nat} (H : 0 < b) (H' : b ∣ a) : - a / b = c ↔ a = b * c := - ⟨Nat.eq_mul_of_div_eq_right H', Nat.div_eq_of_eq_mul_right H⟩ - -protected theorem div_eq_iff_eq_mul_left {a b c : Nat} (H : 0 < b) (H' : b ∣ a) : - a / b = c ↔ a = c * b := by - rw [Nat.mul_comm]; exact Nat.div_eq_iff_eq_mul_right H H' - protected theorem pow_dvd_pow {m n : Nat} (a : Nat) (h : m ≤ n) : a ^ m ∣ a ^ n := by cases Nat.exists_eq_add_of_le h case intro k p => From b39042b32c00b6455153f9df26153680b2dc6d6f Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Fri, 8 Mar 2024 12:37:38 -0800 Subject: [PATCH 06/32] fix: eta-expanded instances at `SynthInstance.lean` (#3638) Remark: this commit removes the `jason1.lean` test. Motivation: It breaks all the time due to changes we make, and it is not clear anymore what it is testing. --------- Co-authored-by: Joachim Breitner --- src/Lean/Meta/SynthInstance.lean | 24 +++++- tests/lean/3057.lean.expected.out | 4 +- tests/lean/815b.lean.expected.out | 4 +- tests/lean/jason1.lean | 50 ----------- tests/lean/jason1.lean.expected.out | 124 ---------------------------- tests/lean/run/instEtaIssue.lean | 24 ++++++ 6 files changed, 50 insertions(+), 180 deletions(-) delete mode 100644 tests/lean/jason1.lean delete mode 100644 tests/lean/jason1.lean.expected.out create mode 100644 tests/lean/run/instEtaIssue.lean diff --git a/src/Lean/Meta/SynthInstance.lean b/src/Lean/Meta/SynthInstance.lean index 3e597da6ff74..40fe335ea2e7 100644 --- a/src/Lean/Meta/SynthInstance.lean +++ b/src/Lean/Meta/SynthInstance.lean @@ -321,6 +321,26 @@ def getSubgoals (lctx : LocalContext) (localInsts : LocalInstances) (xs : Array subgoals := inst.synthOrder.map (mvars[·]!) |>.toList } +/-- +Similar to `mkLambdaFVars`, but ensures result is eta-reduced. +For example, suppose `e` is the local variable `inst x y`, and `xs` is `#[x, y]`, then +the result is `inst` instead of `fun x y => inst x y`. + +We added this auxiliary function because of aliases such as `DecidablePred`. For example, +consider the following definition. +``` +def filter (p : α → Prop) [inst : DecidablePred p] (xs : List α) : List α := + match xs with + | [] => [] + | x :: xs' => if p x then x :: filter p xs' else filter p xs' +``` +Without `mkLambdaFVars'`, the implicit instance at the `filter` applications would be `fun x => inst x` instead of `inst`. +Moreover, the equation lemmas associated with `filter` would have `fun x => inst x` on their right-hand-side. Then, +we would start getting terms such as `fun x => (fun x => inst x) x` when using the equational theorem. +-/ +private def mkLambdaFVars' (xs : Array Expr) (e : Expr) : MetaM Expr := + return (← mkLambdaFVars xs e).eta + /-- Try to synthesize metavariable `mvar` using the instance `inst`. Remark: `mctx` is set using `withMCtx`. @@ -335,7 +355,7 @@ def tryResolve (mvar : Expr) (inst : Instance) : MetaM (Option (MetavarContext withTraceNode `Meta.synthInstance.tryResolve (withMCtx (← getMCtx) do return m!"{exceptOptionEmoji ·} {← instantiateMVars mvarTypeBody} ≟ {← instantiateMVars instTypeBody}") do if (← isDefEq mvarTypeBody instTypeBody) then - let instVal ← mkLambdaFVars xs instVal + let instVal ← mkLambdaFVars' xs instVal if (← isDefEq mvar instVal) then return some ((← getMCtx), subgoals) return none @@ -448,7 +468,7 @@ private def removeUnusedArguments? (mctx : MetavarContext) (mvar : Expr) : MetaM let ys := ys.toArray let mvarType' ← mkForallFVars ys body withLocalDeclD `redf mvarType' fun f => do - let transformer ← mkLambdaFVars #[f] (← mkLambdaFVars xs (mkAppN f ys)) + let transformer ← mkLambdaFVars' #[f] (← mkLambdaFVars' xs (mkAppN f ys)) trace[Meta.synthInstance.unusedArgs] "{mvarType}\nhas unused arguments, reduced type{indentExpr mvarType'}\nTransformer{indentExpr transformer}" return some (mvarType', transformer) diff --git a/tests/lean/3057.lean.expected.out b/tests/lean/3057.lean.expected.out index ccac3ae3097e..34efeb5f83b6 100644 --- a/tests/lean/3057.lean.expected.out +++ b/tests/lean/3057.lean.expected.out @@ -1,7 +1,7 @@ instReprTree instReprListTree -fun a b => instDecidableEqTree a b -fun a b => instDecidableEqListTree a b +instDecidableEqTree +instDecidableEqListTree instBEqTree instBEqListTree instHashableTree diff --git a/tests/lean/815b.lean.expected.out b/tests/lean/815b.lean.expected.out index ebd49e986b72..3b09b50404e5 100644 --- a/tests/lean/815b.lean.expected.out +++ b/tests/lean/815b.lean.expected.out @@ -41,7 +41,7 @@ has unused arguments, reduced type ∀ (b : β), IsSmooth (f b) Transformer - fun redf a b => redf b + fun redf a => redf [Meta.synthInstance] new goal ∀ (b : β), IsSmooth (f b) [Meta.synthInstance.instances] #[inst✝] [Meta.synthInstance] ❌ apply inst✝ to ∀ (b : β), IsSmooth (f b) @@ -74,7 +74,7 @@ has unused arguments, reduced type ∀ (b : β), IsSmooth (f b) Transformer - fun redf a a b => redf b + fun redf a a => redf [Meta.synthInstance] ❌ apply @comp to ∀ (a : α) (a_1 : δ), IsSmooth fun g => f (g a) a_1 [Meta.synthInstance.tryResolve] ❌ IsSmooth fun g => f (g a✝) a ≟ IsSmooth fun a_1 => ?m a✝ a (?m a✝ a a_1) [Meta.synthInstance] ✅ apply @parm to ∀ (a : α) (a_1 : δ), IsSmooth fun g => f (g a) a_1 diff --git a/tests/lean/jason1.lean b/tests/lean/jason1.lean deleted file mode 100644 index 0e1f132ffd0b..000000000000 --- a/tests/lean/jason1.lean +++ /dev/null @@ -1,50 +0,0 @@ -section - variable (G : Type) (T : Type) (Tm : Type) - (EG : G → G → Type) (ET : T → T → Type) (ETm : Tm → Tm → Type) - (getCtx : T → G) (getTy : Tm → T) - inductive CtxSyntaxLayer where - | emp : CtxSyntaxLayer - | snoc : (Γ : G) → (t : T) → EG Γ (getCtx t) → CtxSyntaxLayer -end -section - variable (G : Type) (T : Type) (Tm : Type) - (EG : G → G → Type) (ET : T → T → Type) (ETm : Tm → Tm → Type) - (getCtx : T → G) (getTy : Tm → T) --- : CtxSyntaxLayer G T EG getCtx → G) - - inductive TySyntaxLayer where - | top : {Γ : G} → TySyntaxLayer - | bot : {Γ : G} → TySyntaxLayer - | nat : {Γ : G} → TySyntaxLayer - | arrow : {Γ : G} → (A B : T) → EG Γ (getCtx A) → EG Γ (getCtx B) → TySyntaxLayer - - def getCtxStep : TySyntaxLayer G T EG getCtx → G - | TySyntaxLayer.top (Γ := Γ) .. => Γ - | TySyntaxLayer.bot (Γ := Γ) .. => Γ - | TySyntaxLayer.nat (Γ := Γ) .. => Γ - | TySyntaxLayer.arrow (Γ := Γ) .. => Γ -end -section - variable (G : Type) (T : Type) (Tm : Type) - (EG : G → G → Type) (ET : T → T → Type) (ETm : Tm → Tm → Type) - (EGrfl : ∀ {Γ}, EG Γ Γ) - (getCtx : T → G) (getTy : Tm → T) - (GAlgebra : CtxSyntaxLayer G T EG getCtx → G) (TAlgebra : TySyntaxLayer G T EG getCtx → T) - - inductive TmSyntaxLayer where - | tt : {Γ : G} → TmSyntaxLayer - | zero : {Γ : G} → TmSyntaxLayer - | succ : {Γ : G} → TmSyntaxLayer - | app : {Γ : G} → (A B : T) → (Actx : EG Γ (getCtx A)) → (Bctx : EG Γ (getCtx B)) - → (f x : Tm) - → ET (TAlgebra (TySyntaxLayer.arrow A B Actx Bctx)) (getTy f) - → ET A (getTy x) - → TmSyntaxLayer - - set_option pp.all true - def getTyStep : TmSyntaxLayer G T Tm EG ET getCtx getTy TAlgebra → T - | TmSyntaxLayer.tt .. => TAlgebra TySyntaxLayer.top - | TmSyntaxLayer.zero .. => TAlgebra TySyntaxLayer.nat - | TmSyntaxLayer.succ .. => TAlgebra (TySyntaxLayer.arrow (TAlgebra TySyntaxLayer.nat) (TAlgebra TySyntaxLayer.nat) EGrfl EGrfl) - | TmSyntaxLayer.app (B:=B) .. => B -end diff --git a/tests/lean/jason1.lean.expected.out b/tests/lean/jason1.lean.expected.out deleted file mode 100644 index d211592e8050..000000000000 --- a/tests/lean/jason1.lean.expected.out +++ /dev/null @@ -1,124 +0,0 @@ -jason1.lean:48:119-48:124: error: don't know how to synthesize implicit argument - @EGrfl - (getCtx - (TAlgebra - (@TySyntaxLayer.nat G T EG getCtx - (?m G T Tm EG ET ETm EGrfl getCtx getTy GAlgebra TAlgebra getTyStep Γ✝)))) -context: -G T Tm : Type -EG : G → G → Type -ET : T → T → Type -ETm : Tm → Tm → Type -EGrfl : {Γ : G} → EG Γ Γ -getCtx : T → G -getTy : Tm → T -GAlgebra : CtxSyntaxLayer G T EG getCtx → G -TAlgebra : TySyntaxLayer G T EG getCtx → T -Γ✝ : G -⊢ G -jason1.lean:48:41-48:130: error: don't know how to synthesize implicit argument - @TySyntaxLayer.arrow G T EG getCtx - (getCtx - (TAlgebra - (@TySyntaxLayer.nat G T EG getCtx - (?m G T Tm EG ET ETm EGrfl getCtx getTy GAlgebra TAlgebra getTyStep Γ✝)))) - (TAlgebra - (@TySyntaxLayer.nat G T EG getCtx (?m G T Tm EG ET ETm EGrfl getCtx getTy GAlgebra TAlgebra getTyStep Γ✝))) - (TAlgebra - (@TySyntaxLayer.nat G T EG getCtx (?m G T Tm EG ET ETm EGrfl getCtx getTy GAlgebra TAlgebra getTyStep Γ✝))) - (@EGrfl - (getCtx - (TAlgebra - (@TySyntaxLayer.nat G T EG getCtx - (?m G T Tm EG ET ETm EGrfl getCtx getTy GAlgebra TAlgebra getTyStep Γ✝))))) - (@EGrfl - (getCtx - (TAlgebra - (@TySyntaxLayer.nat G T EG getCtx - (?m G T Tm EG ET ETm EGrfl getCtx getTy GAlgebra TAlgebra getTyStep Γ✝))))) -context: -G T Tm : Type -EG : G → G → Type -ET : T → T → Type -ETm : Tm → Tm → Type -EGrfl : {Γ : G} → EG Γ Γ -getCtx : T → G -getTy : Tm → T -GAlgebra : CtxSyntaxLayer G T EG getCtx → G -TAlgebra : TySyntaxLayer G T EG getCtx → T -Γ✝ : G -⊢ G -jason1.lean:48:100-48:117: error: don't know how to synthesize implicit argument - @TySyntaxLayer.nat G T EG getCtx (?m G T Tm EG ET ETm EGrfl getCtx getTy GAlgebra TAlgebra getTyStep Γ✝) -context: -G T Tm : Type -EG : G → G → Type -ET : T → T → Type -ETm : Tm → Tm → Type -EGrfl : {Γ : G} → EG Γ Γ -getCtx : T → G -getTy : Tm → T -GAlgebra : CtxSyntaxLayer G T EG getCtx → G -TAlgebra : TySyntaxLayer G T EG getCtx → T -Γ✝ : G -⊢ G -jason1.lean:47:40-47:57: error: don't know how to synthesize implicit argument - @TySyntaxLayer.nat G T EG getCtx (?m G T Tm EG ET ETm EGrfl getCtx getTy GAlgebra TAlgebra getTyStep Γ✝) -context: -G T Tm : Type -EG : G → G → Type -ET : T → T → Type -ETm : Tm → Tm → Type -EGrfl : {Γ : G} → EG Γ Γ -getCtx : T → G -getTy : Tm → T -GAlgebra : CtxSyntaxLayer G T EG getCtx → G -TAlgebra : TySyntaxLayer G T EG getCtx → T -Γ✝ : G -⊢ G -jason1.lean:48:71-48:88: error: don't know how to synthesize implicit argument - @TySyntaxLayer.nat G T EG getCtx (?m G T Tm EG ET ETm EGrfl getCtx getTy GAlgebra TAlgebra getTyStep Γ✝) -context: -G T Tm : Type -EG : G → G → Type -ET : T → T → Type -ETm : Tm → Tm → Type -EGrfl : {Γ : G} → EG Γ Γ -getCtx : T → G -getTy : Tm → T -GAlgebra : CtxSyntaxLayer G T EG getCtx → G -TAlgebra : TySyntaxLayer G T EG getCtx → T -Γ✝ : G -⊢ G -jason1.lean:48:125-48:130: error: don't know how to synthesize implicit argument - @EGrfl - (getCtx - (TAlgebra - (@TySyntaxLayer.nat G T EG getCtx - (?m G T Tm EG ET ETm EGrfl getCtx getTy GAlgebra TAlgebra getTyStep Γ✝)))) -context: -G T Tm : Type -EG : G → G → Type -ET : T → T → Type -ETm : Tm → Tm → Type -EGrfl : {Γ : G} → EG Γ Γ -getCtx : T → G -getTy : Tm → T -GAlgebra : CtxSyntaxLayer G T EG getCtx → G -TAlgebra : TySyntaxLayer G T EG getCtx → T -Γ✝ : G -⊢ G -jason1.lean:46:40-46:57: error: don't know how to synthesize implicit argument - @TySyntaxLayer.top G T EG getCtx (?m G T Tm EG ET ETm EGrfl getCtx getTy GAlgebra TAlgebra getTyStep Γ✝) -context: -G T Tm : Type -EG : G → G → Type -ET : T → T → Type -ETm : Tm → Tm → Type -EGrfl : {Γ : G} → EG Γ Γ -getCtx : T → G -getTy : Tm → T -GAlgebra : CtxSyntaxLayer G T EG getCtx → G -TAlgebra : TySyntaxLayer G T EG getCtx → T -Γ✝ : G -⊢ G diff --git a/tests/lean/run/instEtaIssue.lean b/tests/lean/run/instEtaIssue.lean new file mode 100644 index 000000000000..27f18680feb2 --- /dev/null +++ b/tests/lean/run/instEtaIssue.lean @@ -0,0 +1,24 @@ +def filter (p : α → Prop) [DecidablePred p] (xs : List α) : List α := + match xs with + | [] => [] + | x :: xs' => if p x then x :: filter p xs' else filter p xs' + +attribute [simp] filter + +set_option pp.explicit true + +/-- +info: filter._eq_2.{u_1} {α : Type u_1} (p : α → Prop) [@DecidablePred α p] (x : α) (xs' : List α) : + @Eq (List α) (@filter α p inst✝ (@List.cons α x xs')) + (@ite (List α) (p x) (inst✝ x) (@List.cons α x (@filter α p inst✝ xs')) (@filter α p inst✝ xs')) +-/ +#guard_msgs in +#check filter._eq_2 -- We should not have terms of the form `@filter α p (fun x => inst✝ x) xs'` + + +def filter_length (p : α → Prop) [DecidablePred p] : (filter p xs).length ≤ xs.length := by + induction xs with + | nil => simp [filter] + | cons x xs ih => + simp only [filter] + split <;> simp only [List.length] <;> omega From 3acd77a1544c5f6e485bb36dac04c4d5e02e4b4f Mon Sep 17 00:00:00 2001 From: Kyle Miller Date: Sat, 9 Mar 2024 07:30:47 -0800 Subject: [PATCH 07/32] fix: make `elabTermEnsuringType` respect `errToSorry` when there is a type mismatch (#3633) Floris van Doorn [reported on Zulip](https://leanprover.zulipchat.com/#narrow/stream/270676-lean4/topic/have.20tactic.20error.20recovery/near/425283053) that it is confusing that the `have : T := e` tactic completely fails if the body `e` is not of type `T`. This is in contrast to `have : T := by exact e`, which does not completely fail when `e` is not of type `T`. This ends up being caused by `elabTermEnsuringType` throwing an error when it fails to insert a coercion. Now, it detects this case, and it checks the `errToSorry` flag to decide whether to throw the error or to log the error and insert a `sorry`. This is justified by `elabTermEnsuringType` being a frontend to `elabTerm`, which inserts `sorry` on error. An alternative would be to make `ensureType` respect `errToSorry`, but there exists code that expects being able to catch when `ensureType` fails. Making such code manipulate `errToSorry` seems error prone, and this function is not a main entry point to the term elaborator, unlike `elabTermEnsuringType`. --- src/Lean/Elab/Term.lean | 21 ++++++++++++---- tests/lean/issue2260.lean.expected.out | 7 ++++++ .../mulcommErrorMessage.lean.expected.out | 6 +++++ tests/lean/run/haveTactic.lean | 24 +++++++++++++++++++ tests/lean/struct1.lean | 4 ++-- 5 files changed, 56 insertions(+), 6 deletions(-) create mode 100644 tests/lean/run/haveTactic.lean diff --git a/src/Lean/Elab/Term.lean b/src/Lean/Elab/Term.lean index 3fe086dd9c90..c1abb89cb527 100644 --- a/src/Lean/Elab/Term.lean +++ b/src/Lean/Elab/Term.lean @@ -793,10 +793,10 @@ def mkCoe (expectedType : Expr) (e : Expr) (f? : Option Expr := none) (errorMsgH | _ => throwTypeMismatchError errorMsgHeader? expectedType (← inferType e) e f? /-- - If `expectedType?` is `some t`, then ensure `t` and `eType` are definitionally equal. - If they are not, then try coercions. +If `expectedType?` is `some t`, then ensures `t` and `eType` are definitionally equal by inserting a coercion if necessary. - Argument `f?` is used only for generating error messages. -/ +Argument `f?` is used only for generating error messages when inserting coercions fails. +-/ def ensureHasType (expectedType? : Option Expr) (e : Expr) (errorMsgHeader? : Option String := none) (f? : Option Expr := none) : TermElabM Expr := do let some expectedType := expectedType? | return e @@ -1432,9 +1432,22 @@ def addDotCompletionInfo (stx : Syntax) (e : Expr) (expectedType? : Option Expr) def elabTerm (stx : Syntax) (expectedType? : Option Expr) (catchExPostpone := true) (implicitLambda := true) : TermElabM Expr := withRef stx <| elabTermAux expectedType? catchExPostpone implicitLambda stx +/-- +Similar to `Lean.Elab.Term.elabTerm`, but ensures that the type of the elaborated term is `expectedType?` +by inserting coercions if necessary. + +If `errToSorry` is true, then if coercion insertion fails, this function returns `sorry` and logs the error. +Otherwise, it throws the error. +-/ def elabTermEnsuringType (stx : Syntax) (expectedType? : Option Expr) (catchExPostpone := true) (implicitLambda := true) (errorMsgHeader? : Option String := none) : TermElabM Expr := do let e ← elabTerm stx expectedType? catchExPostpone implicitLambda - withRef stx <| ensureHasType expectedType? e errorMsgHeader? + try + withRef stx <| ensureHasType expectedType? e errorMsgHeader? + catch ex => + if (← read).errToSorry && ex matches .error .. then + exceptionToSorry ex expectedType? + else + throw ex /-- Execute `x` and return `some` if no new errors were recorded or exceptions were thrown. Otherwise, return `none`. -/ def commitIfNoErrors? (x : TermElabM α) : TermElabM (Option α) := do diff --git a/tests/lean/issue2260.lean.expected.out b/tests/lean/issue2260.lean.expected.out index 66e493b51cfb..fdcd0e62cb76 100644 --- a/tests/lean/issue2260.lean.expected.out +++ b/tests/lean/issue2260.lean.expected.out @@ -13,3 +13,10 @@ The termination argument has type DNat i : Type but is expected to have type Nat : Type +issue2260.lean:26:15-26:21: error: failed to prove termination, possible solutions: + - Use `have`-expressions to prove the remaining goals + - Use `termination_by` to specify a different well-founded relation + - Use `decreasing_by` to specify your own tactic for discharging this kind of goal +n✝ : Nat +n : DNat n✝ +⊢ sorryAx Nat true < 1 + n✝ + sizeOf n diff --git a/tests/lean/mulcommErrorMessage.lean.expected.out b/tests/lean/mulcommErrorMessage.lean.expected.out index 11c697d680e4..c84c42700008 100644 --- a/tests/lean/mulcommErrorMessage.lean.expected.out +++ b/tests/lean/mulcommErrorMessage.lean.expected.out @@ -14,6 +14,12 @@ has type true = true : Prop but is expected to have type true = false : Prop +mulcommErrorMessage.lean:12:22-12:25: error: type mismatch + rfl +has type + false = false : Prop +but is expected to have type + false = true : Prop mulcommErrorMessage.lean:16:3-17:47: error: type mismatch fun a b => ?m a b has type diff --git a/tests/lean/run/haveTactic.lean b/tests/lean/run/haveTactic.lean new file mode 100644 index 000000000000..54edb02fb6b9 --- /dev/null +++ b/tests/lean/run/haveTactic.lean @@ -0,0 +1,24 @@ +/-! +# Tests for the `have` tactic. +-/ + +/-! +If the body of a `have` fails to elaborate, the tactic completes with a `sorry` for the proof. +-/ +/-- +error: type mismatch + False.elim +has type + False → ?m.6 : Sort ?u.5 +but is expected to have type + True : Prop +--- +info: h : True +⊢ True +-/ +#guard_msgs in +example : True := by + have h : True := + False.elim + trace_state + assumption diff --git a/tests/lean/struct1.lean b/tests/lean/struct1.lean index f6cc3037f4e9..143da10cc036 100644 --- a/tests/lean/struct1.lean +++ b/tests/lean/struct1.lean @@ -31,13 +31,13 @@ structure S := structure S extends A Nat := (x : Nat) -- error -structure S extends A Nat := +structure S' extends A Nat := (x := true) -- error type mismatch structure S extends A Nat := (x : Bool := true) -- error omit type -structure S := +structure S'' := (x : Nat := true) -- error type mismatch private structure S := From 45fccc590622bc9e5a328cb11b734087d10fa49b Mon Sep 17 00:00:00 2001 From: Kyle Miller Date: Sat, 9 Mar 2024 07:31:51 -0800 Subject: [PATCH 08/32] feat: custom eliminators for `induction` and `cases` tactics, and beautiful eliminators for `Nat` (#3629) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Replaces `@[eliminator]` with two attributes `@[induction_eliminator]` and `@[cases_eliminator]` for defining custom eliminators for the `induction` and `cases` tactics, respectively. Adds `Nat.recAux` and `Nat.casesAuxOn`, which are eliminators that are defeq to `Nat.rec` and `Nat.casesOn`, but these use `0` and `n + 1` rather than `Nat.zero` and `Nat.succ n`. For example, using `induction` to prove that the factorial function is positive now has the following goal states (thanks also to #3616 for the goal state after unfolding). ```lean example : 0 < fact x := by induction x with | zero => decide | succ x ih => /- x : Nat ih : 0 < fact x ⊢ 0 < fact (x + 1) -/ unfold fact /- ... ⊢ 0 < (x + 1) * fact x -/ simpa using ih ``` Thanks to @adamtopaz for initial work on splitting the `@[eliminator]` attribute. --- src/Init/Data/Array/Basic.lean | 2 +- src/Init/Data/Nat/Basic.lean | 27 +- src/Init/Data/Nat/Div.lean | 6 +- src/Lean/Elab/Tactic/Induction.lean | 4 +- src/Lean/Meta/Tactic/ElimInfo.lean | 40 +- stage0/stdlib/Init/Data/Nat/Basic.c | 117 + stage0/stdlib/Lean/Elab/Tactic/Induction.c | 172 +- stage0/stdlib/Lean/Meta/Tactic/ElimInfo.c | 5051 ++++++++++------- stage0/stdlib/Lean/Meta/Tactic/Simp/Main.c | 376 +- tests/lean/inductionErrors.lean.expected.out | 4 +- .../interactive/anonHyp.lean.expected.out | 4 +- .../interactive/goalIssue.lean.expected.out | 4 +- .../lean/interactive/hover.lean.expected.out | 15 +- .../interactive/infoIssues.lean.expected.out | 14 +- .../interactive/plainGoal.lean.expected.out | 32 +- tests/lean/run/1337.lean | 2 +- tests/lean/run/aStructPerfIssue.lean | 6 +- tests/lean/run/casesAnyTypeIssue.lean | 2 +- tests/lean/run/compatibleTypesEtaIssue.lean | 2 - tests/lean/run/defaultEliminator.lean | 2 +- tests/lean/run/eliminatorImplicitTargets.lean | 2 +- tests/lean/run/internalizeCasesIssue.lean | 2 +- tests/lean/run/lcnfCastIssue.lean | 4 +- tests/lean/run/lcnfInliningIssue.lean | 2 +- tests/lean/simp_trace.lean.expected.out | 2 +- tests/lean/unfold1.lean.expected.out | 2 +- tests/lean/unsolvedIndCases.lean.expected.out | 10 +- 27 files changed, 3694 insertions(+), 2212 deletions(-) diff --git a/src/Init/Data/Array/Basic.lean b/src/Init/Data/Array/Basic.lean index 5839247da318..7a1556efa513 100644 --- a/src/Init/Data/Array/Basic.lean +++ b/src/Init/Data/Array/Basic.lean @@ -809,7 +809,7 @@ where rfl go (i : Nat) (hi : i ≤ as.size) : toListLitAux as n hsz i hi (as.data.drop i) = as.data := by - cases i <;> simp [getLit_eq, List.get_drop_eq_drop, toListLitAux, List.drop, go] + induction i <;> simp [getLit_eq, List.get_drop_eq_drop, toListLitAux, List.drop, *] def isPrefixOfAux [BEq α] (as bs : Array α) (hle : as.size ≤ bs.size) (i : Nat) : Bool := if h : i < as.size then diff --git a/src/Init/Data/Nat/Basic.lean b/src/Init/Data/Nat/Basic.lean index c9287456da85..49d4f35f9f17 100644 --- a/src/Init/Data/Nat/Basic.lean +++ b/src/Init/Data/Nat/Basic.lean @@ -10,6 +10,29 @@ universe u namespace Nat +/-- Compiled version of `Nat.rec` so that we can define `Nat.recAux` to be defeq to `Nat.rec`. +This is working around the fact that the compiler does not currently support recursors. -/ +private def recCompiled {motive : Nat → Sort u} (zero : motive zero) (succ : (n : Nat) → motive n → motive (Nat.succ n)) : (t : Nat) → motive t + | .zero => zero + | .succ n => succ n (recCompiled zero succ n) + +@[csimp] +private theorem rec_eq_recCompiled : @Nat.rec = @Nat.recCompiled := + funext fun _ => funext fun _ => funext fun succ => funext fun t => + Nat.recOn t rfl (fun n ih => congrArg (succ n) ih) + +/-- Recursor identical to `Nat.rec` but uses notations `0` for `Nat.zero` and `· + 1` for `Nat.succ`. +Used as the default `Nat` eliminator by the `induction` tactic. -/ +@[elab_as_elim, induction_eliminator] +protected abbrev recAux {motive : Nat → Sort u} (zero : motive 0) (succ : (n : Nat) → motive n → motive (n + 1)) (t : Nat) : motive t := + Nat.rec zero succ t + +/-- Recursor identical to `Nat.casesOn` but uses notations `0` for `Nat.zero` and `· + 1` for `Nat.succ`. +Used as the default `Nat` eliminator by the `cases` tactic. -/ +@[elab_as_elim, cases_eliminator] +protected abbrev casesAuxOn {motive : Nat → Sort u} (t : Nat) (zero : motive 0) (succ : (n : Nat) → motive (n + 1)) : motive t := + Nat.casesOn t zero succ + /-- `Nat.fold` evaluates `f` on the numbers up to `n` exclusive, in increasing order: * `Nat.fold f 3 init = init |> f 0 |> f 1 |> f 2` @@ -738,7 +761,7 @@ theorem zero_lt_sub_of_lt (h : i < a) : 0 < a - i := by | zero => contradiction | succ a ih => match Nat.eq_or_lt_of_le h with - | Or.inl h => injection h with h; subst h; rw [←Nat.add_one, Nat.add_sub_self_left]; decide + | Or.inl h => injection h with h; subst h; rw [Nat.add_sub_self_left]; decide | Or.inr h => have : 0 < a - i := ih (Nat.lt_of_succ_lt_succ h) exact Nat.lt_of_lt_of_le this (Nat.sub_le_succ_sub _ _) @@ -883,7 +906,7 @@ protected theorem sub_pos_of_lt (h : m < n) : 0 < n - m := protected theorem sub_sub (n m k : Nat) : n - m - k = n - (m + k) := by induction k with | zero => simp - | succ k ih => rw [Nat.add_succ, Nat.sub_succ, Nat.sub_succ, ih] + | succ k ih => rw [Nat.add_succ, Nat.sub_succ, Nat.add_succ, Nat.sub_succ, ih] protected theorem sub_le_sub_left (h : n ≤ m) (k : Nat) : k - m ≤ k - n := match m, le.dest h with diff --git a/src/Init/Data/Nat/Div.lean b/src/Init/Data/Nat/Div.lean index fbdfb23d1629..e51da62f2c6a 100644 --- a/src/Init/Data/Nat/Div.lean +++ b/src/Init/Data/Nat/Div.lean @@ -198,11 +198,11 @@ theorem le_div_iff_mul_le (k0 : 0 < k) : x ≤ y / k ↔ x * k ≤ y := by induction y, k using mod.inductionOn generalizing x with (rw [div_eq]; simp [h]; cases x with | zero => simp [zero_le] | succ x => ?_) | base y k h => - simp [not_succ_le_zero x, succ_mul, Nat.add_comm] - refine Nat.lt_of_lt_of_le ?_ (Nat.le_add_right ..) + simp only [add_one, succ_mul, false_iff, Nat.not_le] + refine Nat.lt_of_lt_of_le ?_ (Nat.le_add_left ..) exact Nat.not_le.1 fun h' => h ⟨k0, h'⟩ | ind y k h IH => - rw [← add_one, Nat.add_le_add_iff_right, IH k0, succ_mul, + rw [Nat.add_le_add_iff_right, IH k0, succ_mul, ← Nat.add_sub_cancel (x*k) k, Nat.sub_le_sub_iff_right h.2, Nat.add_sub_cancel] protected theorem div_div_eq_div_mul (m n k : Nat) : m / n / k = m / (n * k) := by diff --git a/src/Lean/Elab/Tactic/Induction.lean b/src/Lean/Elab/Tactic/Induction.lean index 10a63a5f9022..6852e48c69b6 100644 --- a/src/Lean/Elab/Tactic/Induction.lean +++ b/src/Lean/Elab/Tactic/Induction.lean @@ -534,9 +534,9 @@ private def elabTermForElim (stx : Syntax) : TermElabM Expr := do return e -- `optElimId` is of the form `("using" term)?` -private def getElimNameInfo (optElimId : Syntax) (targets : Array Expr) (induction : Bool): TacticM ElimInfo := do +private def getElimNameInfo (optElimId : Syntax) (targets : Array Expr) (induction : Bool) : TacticM ElimInfo := do if optElimId.isNone then - if let some elimName ← getCustomEliminator? targets then + if let some elimName ← getCustomEliminator? targets induction then return ← getElimInfo elimName unless targets.size == 1 do throwError "eliminator must be provided when multiple targets are used (use 'using '), and no default eliminator has been registered using attribute `[eliminator]`" diff --git a/src/Lean/Meta/Tactic/ElimInfo.lean b/src/Lean/Meta/Tactic/ElimInfo.lean index e1b6ab259028..7e954798f6a8 100644 --- a/src/Lean/Meta/Tactic/ElimInfo.lean +++ b/src/Lean/Meta/Tactic/ElimInfo.lean @@ -123,17 +123,18 @@ where return (implicits, targets') structure CustomEliminator where + induction : Bool typeNames : Array Name elimName : Name -- NB: Do not store the ElimInfo, it can contain MVars deriving Inhabited, Repr structure CustomEliminators where - map : SMap (Array Name) Name := {} + map : SMap (Bool × Array Name) Name := {} deriving Inhabited, Repr def addCustomEliminatorEntry (es : CustomEliminators) (e : CustomEliminator) : CustomEliminators := match es with - | { map := map } => { map := map.insert e.typeNames e.elimName } + | { map := map } => { map := map.insert (e.induction, e.typeNames) e.elimName } builtin_initialize customEliminatorExt : SimpleScopedEnvExtension CustomEliminator CustomEliminators ← registerSimpleScopedEnvExtension { @@ -142,7 +143,7 @@ builtin_initialize customEliminatorExt : SimpleScopedEnvExtension CustomEliminat finalizeImport := fun { map := map } => { map := map.switch } } -def mkCustomEliminator (elimName : Name) : MetaM CustomEliminator := do +def mkCustomEliminator (elimName : Name) (induction : Bool) : MetaM CustomEliminator := do let elimInfo ← getElimInfo elimName let info ← getConstInfo elimName forallTelescopeReducing info.type fun xs _ => do @@ -164,29 +165,46 @@ def mkCustomEliminator (elimName : Name) : MetaM CustomEliminator := do let xType ← inferType x let .const typeName .. := xType.getAppFn | throwError "unexpected eliminator target type{indentExpr xType}" typeNames := typeNames.push typeName - return { typeNames, elimName} + return { induction, typeNames, elimName } -def addCustomEliminator (declName : Name) (attrKind : AttributeKind) : MetaM Unit := do - let e ← mkCustomEliminator declName +def addCustomEliminator (declName : Name) (attrKind : AttributeKind) (induction : Bool) : MetaM Unit := do + let e ← mkCustomEliminator declName induction customEliminatorExt.add e attrKind builtin_initialize registerBuiltinAttribute { - name := `eliminator - descr := "custom eliminator for `cases` and `induction` tactics" + name := `induction_eliminator + descr := "custom `rec`-like eliminator for the `induction` tactic" add := fun declName _ attrKind => do - discard <| addCustomEliminator declName attrKind |>.run {} {} + discard <| addCustomEliminator declName attrKind (induction := true) |>.run {} {} } +builtin_initialize + registerBuiltinAttribute { + name := `cases_eliminator + descr := "custom `casesOn`-like eliminator for the `cases` tactic" + add := fun declName _ attrKind => do + discard <| addCustomEliminator declName attrKind (induction := false) |>.run {} {} + } + +/-- Gets all the eliminators defined using the `@[induction_eliminator]` and `@[cases_eliminator]` attributes. -/ def getCustomEliminators : CoreM CustomEliminators := do return customEliminatorExt.getState (← getEnv) -def getCustomEliminator? (targets : Array Expr) : MetaM (Option Name) := do +/-- +Gets an eliminator appropriate for the provided array of targets. +If `induction` is `true` then returns a matching eliminator defined using the `@[induction_eliminator]` attribute +and otherwise returns one defined using the `@[cases_eliminator]` attribute. + +The `@[induction_eliminator]` attribute is for the `induction` tactic +and the `@[cases_eliminator]` attribute is for the `cases` tactic. +-/ +def getCustomEliminator? (targets : Array Expr) (induction : Bool) : MetaM (Option Name) := do let mut key := #[] for target in targets do let targetType := (← instantiateMVars (← inferType target)).headBeta let .const declName .. := targetType.getAppFn | return none key := key.push declName - return customEliminatorExt.getState (← getEnv) |>.map.find? key + return customEliminatorExt.getState (← getEnv) |>.map.find? (induction, key) end Lean.Meta diff --git a/stage0/stdlib/Init/Data/Nat/Basic.c b/stage0/stdlib/Init/Data/Nat/Basic.c index 23a6dc8bbdba..db36e046c749 100644 --- a/stage0/stdlib/Init/Data/Nat/Basic.c +++ b/stage0/stdlib/Init/Data/Nat/Basic.c @@ -18,9 +18,11 @@ LEAN_EXPORT lean_object* l_Nat_foldTR___rarg(lean_object*, lean_object*, lean_ob LEAN_EXPORT uint8_t l_Nat_blt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_repeatTR(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Basic_0__Nat_any_match__1_splitter(lean_object*); +LEAN_EXPORT lean_object* l_Nat_casesAuxOn___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Basic_0__Nat_beq_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Basic_0__Nat_fold_match__1_splitter(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_min___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Nat_recAux(lean_object*); LEAN_EXPORT lean_object* l_Nat_foldRev(lean_object*); LEAN_EXPORT lean_object* l_Nat_allTR_loop___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_instTransNatLtInstLTNatLeInstLENat; @@ -32,12 +34,15 @@ LEAN_EXPORT lean_object* l_Nat_instMaxNat(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_repeatTR_loop(lean_object*); LEAN_EXPORT lean_object* l_Nat_fold___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_max(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Basic_0__Nat_recCompiled___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_blt___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Prod_foldI___rarg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Nat_recAux___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Basic_0__Nat_beq_match__1_splitter___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_all___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_repeatTR_loop___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_instMaxNat___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Basic_0__Nat_recCompiled___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Prod_allI(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_instTransNatLeInstLENat; LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Basic_0__Nat_any_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*); @@ -48,11 +53,13 @@ LEAN_EXPORT uint8_t l_Nat_anyTR_loop(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Nat_allTR(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_max___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Prod_allI___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Nat_casesAuxOn(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Basic_0__Nat_fold_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Nat_anyTR(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_foldTR_loop___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_instAntisymmNatNotLtInstLTNat; LEAN_EXPORT uint8_t l_Nat_allTR_loop(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Nat_casesAuxOn___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_foldTR_loop(lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Prod_anyI(lean_object*, lean_object*); @@ -67,8 +74,10 @@ LEAN_EXPORT lean_object* l_Nat_instAntisymmNatLeInstLENat; LEAN_EXPORT lean_object* l_Nat_foldRev___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Basic_0__Nat_beq_match__1_splitter(lean_object*); LEAN_EXPORT lean_object* l_Nat_min(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Nat_recAux___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_any(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_fold(lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Basic_0__Nat_recCompiled(lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_allTR___boxed(lean_object*, lean_object*); @@ -77,6 +86,114 @@ LEAN_EXPORT lean_object* l_Prod_foldI___rarg(lean_object*, lean_object*, lean_ob LEAN_EXPORT lean_object* l_Nat_foldTR(lean_object*); LEAN_EXPORT lean_object* l_Nat_any___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_anyTR___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Basic_0__Nat_recCompiled___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_unsigned_to_nat(0u); +x_5 = lean_nat_dec_eq(x_3, x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_6 = lean_unsigned_to_nat(1u); +x_7 = lean_nat_sub(x_3, x_6); +lean_inc(x_2); +x_8 = l___private_Init_Data_Nat_Basic_0__Nat_recCompiled___rarg(x_1, x_2, x_7); +x_9 = lean_apply_2(x_2, x_7, x_8); +return x_9; +} +else +{ +lean_dec(x_2); +lean_inc(x_1); +return x_1; +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Basic_0__Nat_recCompiled(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Init_Data_Nat_Basic_0__Nat_recCompiled___rarg___boxed), 3, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Basic_0__Nat_recCompiled___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Init_Data_Nat_Basic_0__Nat_recCompiled___rarg(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Nat_recAux___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Init_Data_Nat_Basic_0__Nat_recCompiled___rarg(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Nat_recAux(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Nat_recAux___rarg___boxed), 3, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Nat_recAux___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Nat_recAux___rarg(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Nat_casesAuxOn___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_unsigned_to_nat(0u); +x_5 = lean_nat_dec_eq(x_1, x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_6 = lean_unsigned_to_nat(1u); +x_7 = lean_nat_sub(x_1, x_6); +x_8 = lean_apply_1(x_3, x_7); +return x_8; +} +else +{ +lean_dec(x_3); +lean_inc(x_2); +return x_2; +} +} +} +LEAN_EXPORT lean_object* l_Nat_casesAuxOn(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Nat_casesAuxOn___rarg___boxed), 3, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Nat_casesAuxOn___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Nat_casesAuxOn___rarg(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} LEAN_EXPORT lean_object* l_Nat_fold___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { diff --git a/stage0/stdlib/Lean/Elab/Tactic/Induction.c b/stage0/stdlib/Lean/Elab/Tactic/Induction.c index d0553e6b225f..801124782169 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Induction.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Induction.c @@ -26,7 +26,7 @@ lean_object* l___private_Init_Util_0__outOfBounds___rarg(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalInduction_declRange___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getAltName___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_expandCases_x3f(lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__3; +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__1; lean_object* l_Lean_Meta_getElimExprInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_go___spec__9___lambda__6___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getElimNameInfo___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*); @@ -42,7 +42,6 @@ static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_evalInduction 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_object* l_Lean_Meta_throwTacticEx___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__2; lean_object* l_Lean_Elab_CommandContextInfo_save___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSavedPartialInfoContext___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_saveAltVarsInfo___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_getArgExpectedType(lean_object*); @@ -67,6 +66,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tac lean_object* lean_whnf(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_expandInductionAlts_x3f(lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getElimNameInfo___lambda__5___closed__1; +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__11; LEAN_EXPORT uint8_t l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_isMultiAlt(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_generalizeVars(lean_object*, 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_synthesizeSyntheticMVars(uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -95,6 +95,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalCases_declRange___closed 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___regBuiltin_Lean_Elab_Tactic_evalCases(lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__9; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_expandMultiAlt_x3f___spec__1___closed__1; static lean_object* l_Lean_Elab_Tactic_elabCasesTargets___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_generalizeTargets___spec__1___boxed(lean_object*, lean_object*, lean_object*); @@ -131,7 +132,6 @@ static lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_g static lean_object* l_Lean_Elab_Tactic_elabCasesTargets___closed__1; 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*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_saveAltVarsInfo___spec__1(lean_object*, uint8_t, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__7; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_checkAltsOfOptInductionAlts___spec__1___closed__1; 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*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_elabTermForElim___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -154,7 +154,6 @@ lean_object* l_Lean_MVarId_setKind(lean_object*, uint8_t, lean_object*, lean_obj LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_Result_alts___default; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts___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_Tactic_Induction_0__Lean_Elab_Tactic_getAltRHS(lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__16; lean_object* l_Lean_addBuiltinDeclarationRanges(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getElimNameInfo___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___regBuiltin_Lean_Elab_Tactic_evalInduction_declRange(lean_object*); @@ -166,6 +165,7 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_evalIndu LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getElimNameInfo___lambda__3(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___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_addNewArg(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_LocalContext_empty; +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__6; uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_expandMultiAlt_x3f(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_go___spec__9___lambda__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*); @@ -201,6 +201,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_ LEAN_EXPORT uint8_t l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_altHasExplicitModifier(lean_object*); LEAN_EXPORT uint8_t l_Lean_Elab_Tactic_isHoleRHS(lean_object*); lean_object* l_Lean_Syntax_getNumArgs(lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__4; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_elabCasesTargets___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_Elab_Tactic_Induction_0__Lean_Elab_Tactic_generalizeTargets___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___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getAltName___closed__5; @@ -222,6 +223,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tac LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getAltNameStx(lean_object*); static lean_object* l_Lean_Elab_Tactic_ElimApp_setMotiveArg___closed__4; lean_object* l_Lean_Elab_Term_withoutHeedElabAsElimImp___rarg(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_Induction___hyg_10488____closed__16; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_ElimApp_evalAlts_go___spec__6(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___regBuiltin_Lean_Elab_Tactic_evalInduction___closed__3; @@ -231,7 +233,6 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tacti 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*); static lean_object* l_Lean_Elab_Tactic_ElimApp_setMotiveArg___closed__2; LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_Tactic_ElimApp_evalAlts_go___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__12; lean_object* l_Lean_Elab_Tactic_elabTerm(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_ElimApp_instInhabitedAlt___closed__1; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_generalizeTargets___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -239,6 +240,7 @@ static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_evalInduction LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_ElimApp_reorderAlts___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_saveAltVarsInfo(lean_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_elabCasesTargets___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_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__10; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_go___spec__9___lambda__6___boxed__const__1; lean_object* lean_st_ref_take(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getAltNameStx___boxed(lean_object*); @@ -249,7 +251,6 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_evalIndu LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_go___spec__9___lambda__9___boxed(lean_object**); static lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getElimNameInfo___lambda__4___closed__2; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_go___spec__9___lambda__8___closed__6; -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__13; static lean_object* l_Lean_Elab_withSaveInfoContext___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_saveAltVarsInfo___spec__2___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_evalInduction_checkTargets___spec__1___at_Lean_Elab_Tactic_evalInduction_checkTargets___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_Elab_Tactic_ElimApp_evalAlts___boxed(lean_object**); @@ -289,7 +290,6 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_ElimApp static lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getAltName___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalCases___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_go___spec__9___lambda__11___boxed(lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__9; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_generalizeTargets___boxed__const__1; 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*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalInduction___closed__6; @@ -300,13 +300,15 @@ lean_object* l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0_ LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getAltVars(lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_checkAltNames___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_go___spec__9___closed__1; -lean_object* l_Lean_Meta_getCustomEliminator_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__15; +lean_object* l_Lean_Meta_getCustomEliminator_x3f(lean_object*, uint8_t, 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_EXPORT lean_object* l_panic___at_Lean_Elab_Tactic_elabCasesTargets___spec__3(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*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_evalInduction_checkTargets___spec__1___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_ElimApp_evalAlts_go___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*); 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*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__5; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_saveAltVarsInfo___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Tactic_elabCasesTargets___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_Lean_Meta_whnfForall(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -315,7 +317,6 @@ static lean_object* l_Lean_Elab_Tactic_elabCasesTargets___lambda__1___closed__5; lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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*); uint8_t l_List_isEmpty___rarg(lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__6; lean_object* lean_st_mk_ref(lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_synthesizeInstMVarCore(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_Induction_0__Lean_Elab_Tactic_getAltsOfOptInductionAlts___boxed(lean_object*); @@ -339,12 +340,12 @@ lean_object* l_Lean_Meta_addImplicitTargets(lean_object*, lean_object*, lean_obj lean_object* l_Lean_Expr_constName_x3f(lean_object*); lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_addDotCompletionInfo___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_getAltNumFields___closed__1; -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalCases___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_go___spec__9___lambda__2___closed__2; static lean_object* l_Lean_Elab_Tactic_ElimApp_mkElimApp___closed__2; static lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_go___lambda__2___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabCasesTargets___boxed__const__1; +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__12; static lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getUserGeneralizingFVarIds___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_generalizeTargets___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_Tactic_ElimApp_evalAlts_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*); @@ -355,6 +356,7 @@ lean_object* l_Lean_Syntax_getSepArgs(lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_go___spec__9___lambda__8___closed__7; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_go___boxed(lean_object**); static lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getElimNameInfo___lambda__3___closed__3; +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__7; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_checkAltsOfOptInductionAlts___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_go___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalInduction___closed__1; @@ -396,7 +398,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSa LEAN_EXPORT lean_object* l_Lean_Elab_withSaveInfoContext___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_saveAltVarsInfo___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_getAltNumFields___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_instInhabitedMetaM___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__10; extern lean_object* l_Lean_instInhabitedFVarId; static lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getElimNameInfo___lambda__3___closed__1; static lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_generalizeTargets___closed__1; @@ -404,7 +405,6 @@ lean_object* l_Array_append___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_ElimApp_evalAlts_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_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofExpr(lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__15; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getElimNameInfo___lambda__5(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_ElimApp_mkElimApp_loop___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalCases___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*); @@ -459,6 +459,7 @@ uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_go___spec__9___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*, 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_object* l_Lean_MVarId_getDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_withDeclName___spec__3___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__13; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getAltDArrow(lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_evalInduction___lambda__2___closed__2; @@ -480,15 +481,14 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Tact lean_object* l_Lean_Meta_getFVarSetToGeneralize(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_ElimApp_setMotiveArg___closed__6; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_go___spec__9___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_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487_(lean_object*); 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*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getUserGeneralizingFVarIds___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_Induction___hyg_10488____closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getInductiveValFromMajor___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_mkHasTypeButIsExpectedMsg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_generalizeVars___spec__1___closed__2; 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_object* l_Lean_RBTree_toArray___at_Lean_Meta_getFVarsToGeneralize___spec__1(lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__8; 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_evalInduction___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* 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*); @@ -512,8 +512,8 @@ static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalA LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_checkAltsOfOptInductionAlts___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_ElimApp_evalAlts___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_go___spec__9___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_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488_(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalInduction_declRange___closed__6; -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__14; lean_object* l_Array_ofSubarray___rarg(lean_object*); static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_checkAltNames___spec__2___lambda__1___closed__1; uint8_t l_Lean_LocalDecl_binderInfo(lean_object*); @@ -558,6 +558,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tac LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_checkAltNames(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_go___spec__9___lambda__8___closed__5; lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__2; lean_object* l_Lean_Expr_bindingBody_x21(lean_object*); lean_object* l_List_redLength___rarg(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_expandInduction_x3f(lean_object*); @@ -568,7 +569,7 @@ LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at___private_Lean_Elab_Ta LEAN_EXPORT lean_object* l_Lean_withoutModifyingState___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_getNumExplicitFields___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Tactic_evalTactic(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_withTacticInfoContext___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_Induction___hyg_10487____closed__4; +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__14; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_go___spec__7(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_go___spec__9___lambda__8___closed__4; lean_object* lean_string_append(lean_object*, lean_object*); @@ -577,6 +578,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalInduction_checkTargets(lean_obje LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_go___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*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_shouldGeneralizeTarget___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_Induction___hyg_10488____closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getFirstAltLhs___boxed(lean_object*); lean_object* lean_array_get_size(lean_object*); static lean_object* l_Lean_Elab_Tactic_evalCases___lambda__2___closed__1; @@ -591,7 +593,6 @@ uint8_t lean_nat_dec_le(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalCases_declRange___closed__6; 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*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__11; lean_object* l_Lean_Elab_admitGoal(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Cases_unifyEqs_x3f(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___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_saveAltVarsInfo___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -638,7 +639,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalAlt___lambda__4(lean_object*, le LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_State_targetPos___default; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_getFType___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_ElimApp_reorderAlts___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__1; uint8_t l_Array_isEmpty___rarg(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalInduction_declRange___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getFirstAltLhs(lean_object* x_1) { @@ -16851,7 +16851,7 @@ lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_2); -x_63 = l_Lean_Meta_getCustomEliminator_x3f(x_2, x_8, x_9, x_10, x_11, x_12); +x_63 = l_Lean_Meta_getCustomEliminator_x3f(x_2, x_3, x_8, x_9, x_10, x_11, x_12); if (lean_obj_tag(x_63) == 0) { lean_object* x_64; @@ -22306,7 +22306,7 @@ x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__1() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -22316,7 +22316,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_Induction___hyg_10487____closed__2() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -22326,27 +22326,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_Induction___hyg_10487____closed__3() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____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_Induction___hyg_10487____closed__2; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__2; x_2 = l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getUserGeneralizingFVarIds___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_Induction___hyg_10487____closed__4() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____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_Induction___hyg_10487____closed__3; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__3; x_2 = l___regBuiltin_Lean_Elab_Tactic_evalInduction___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_Induction___hyg_10487____closed__5() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__5() { _start: { lean_object* x_1; @@ -22354,17 +22354,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_Induction___hyg_10487____closed__6() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____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_Induction___hyg_10487____closed__4; -x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__5; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__4; +x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____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_Induction___hyg_10487____closed__7() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__7() { _start: { lean_object* x_1; @@ -22372,47 +22372,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_Induction___hyg_10487____closed__8() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____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_Induction___hyg_10487____closed__6; -x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__7; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__6; +x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____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_Induction___hyg_10487____closed__9() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____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_Induction___hyg_10487____closed__8; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__8; x_2 = l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getAltName___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_Induction___hyg_10487____closed__10() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____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_Induction___hyg_10487____closed__9; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__9; x_2 = l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getUserGeneralizingFVarIds___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_Induction___hyg_10487____closed__11() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____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_Induction___hyg_10487____closed__10; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__10; x_2 = l___regBuiltin_Lean_Elab_Tactic_evalInduction___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_Induction___hyg_10487____closed__12() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__12() { _start: { lean_object* x_1; @@ -22420,17 +22420,17 @@ x_1 = lean_mk_string_from_bytes("Induction", 9); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__13() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____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_Induction___hyg_10487____closed__11; -x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__12; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__11; +x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____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_Induction___hyg_10487____closed__14() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__14() { _start: { lean_object* x_1; @@ -22438,33 +22438,33 @@ 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_Induction___hyg_10487____closed__15() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____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_Induction___hyg_10487____closed__13; -x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__14; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__13; +x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____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_Induction___hyg_10487____closed__16() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____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_Induction___hyg_10487____closed__15; -x_2 = lean_unsigned_to_nat(10487u); +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__15; +x_2 = lean_unsigned_to_nat(10488u); 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_Induction___hyg_10487_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488_(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_Induction___hyg_10487____closed__1; +x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__1; x_3 = 0; -x_4 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__16; +x_4 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__16; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); if (lean_obj_tag(x_5) == 0) { @@ -22843,39 +22843,39 @@ lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_evalCases_declRange___close if (builtin) {res = l___regBuiltin_Lean_Elab_Tactic_evalCases_declRange(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_Induction___hyg_10487____closed__1 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__1(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__1); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__2 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__2(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__2); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__3 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__3(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__3); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__4 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__4(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__4); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__5 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__5(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__5); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__6 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__6(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__6); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__7 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__7(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__7); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__8 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__8(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__8); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__9 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__9(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__9); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__10 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__10(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__10); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__11 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__11(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__11); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__12 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__12(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__12); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__13 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__13(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__13); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__14 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__14(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__14); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__15 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__15(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__15); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__16 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__16(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487____closed__16); -if (builtin) {res = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10487_(lean_io_mk_world()); +}l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__1 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__1); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__2 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__2); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__3 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__3(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__3); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__4 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__4(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__4); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__5 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__5(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__5); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__6 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__6(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__6); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__7 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__7(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__7); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__8 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__8(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__8); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__9 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__9(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__9); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__10 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__10(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__10); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__11 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__11(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__11); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__12 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__12(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__12); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__13 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__13(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__13); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__14 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__14(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__14); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__15 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__15(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__15); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__16 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__16(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488____closed__16); +if (builtin) {res = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_10488_(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/ElimInfo.c b/stage0/stdlib/Lean/Meta/Tactic/ElimInfo.c index 31b8f4f1e871..0b15aeda6bba 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/ElimInfo.c +++ b/stage0/stdlib/Lean/Meta/Tactic/ElimInfo.c @@ -14,114 +14,130 @@ extern "C" { #endif LEAN_EXPORT lean_object* l_Lean_Meta_getElimInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__1; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Meta_mkCustomEliminator___spec__19___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8___boxed(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_2270____spec__11___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_addCustomEliminatorEntry(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__3___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_getCustomEliminator_x3f___spec__4___boxed(lean_object*, lean_object*, lean_object*); lean_object* l___private_Init_Util_0__outOfBounds___rarg(lean_object*); static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__5; lean_object* l_Lean_Name_reprPrec(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getElimExprInfo(lean_object*, lean_object*, 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_205____closed__3; -LEAN_EXPORT lean_object* l_Lean_SMap_fold___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__2(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at_Lean_Meta_addImplicitTargets___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Meta_getCustomEliminator_x3f___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMap_insert___at_Lean_Meta_addCustomEliminatorEntry___spec__9(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_205____closed__22; static lean_object* l_Lean_Meta_getElimExprInfo___lambda__2___closed__1; -LEAN_EXPORT lean_object* l_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____spec__1(lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__8; 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_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8___closed__6; static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__18; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157_(lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); size_t lean_usize_shift_right(size_t, size_t); static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addCustomEliminator___spec__1___closed__16; static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__13; LEAN_EXPORT lean_object* l_Lean_Meta_CustomEliminators_map___default; static lean_object* l_Std_Range_forIn_loop___at_Lean_Meta_mkCustomEliminator___spec__19___closed__2; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__1; static lean_object* l_Lean_Meta_getElimExprInfo___closed__7; LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at_Lean_Meta_mkCustomEliminator___spec__6(lean_object*, lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__25; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_mkCustomEliminator___spec__9(lean_object*, lean_object*, size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_shouldVisit(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getElimExprInfo___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Option_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____spec__1___closed__2; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_addImplicitTargets___spec__2___closed__2; static lean_object* l_Lean_Meta_addImplicitTargets_collect___lambda__2___closed__4; -static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__10; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_Meta_mkCustomEliminator___spec__10___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_addImplicitTargets___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_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__7; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__3; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__4; +static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____closed__7; static lean_object* l_Lean_Meta_addImplicitTargets_collect___lambda__2___closed__1; lean_object* l_Lean_ConstantInfo_type(lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMapImp_find_x3f___at_Lean_Meta_getCustomEliminator_x3f___spec__8(lean_object*, lean_object*); lean_object* l_Lean_MetavarContext_getDecl(lean_object*, lean_object*); static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addCustomEliminator___spec__1___closed__7; uint8_t lean_usize_dec_le(size_t, size_t); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__13; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__11; LEAN_EXPORT lean_object* l_Lean_Meta_instInhabitedCustomEliminators; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__22; uint64_t lean_uint64_mix_hash(uint64_t, uint64_t); static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addCustomEliminator___spec__1___closed__11; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____closed__2; lean_object* l_Lean_Meta_isExprDefEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_uint64_to_usize(uint64_t); LEAN_EXPORT lean_object* l_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____spec__2(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070_(lean_object*); LEAN_EXPORT uint8_t l_Array_contains___at_Lean_Meta_getElimExprInfo___spec__4(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__2___closed__1; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__3; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_addImplicitTargets___spec__3(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946_(lean_object*); +static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_instReprCustomEliminators; -static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__2; +LEAN_EXPORT lean_object* l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8___boxed(lean_object*, lean_object*); lean_object* l_Lean_Expr_sort___override(lean_object*); static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__24; static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__11; -LEAN_EXPORT lean_object* l_Lean_Meta_addCustomEliminator___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_addCustomEliminator___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__26; lean_object* l_Lean_mkHashSetImp___rarg(lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__16; 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_addImplicitTargets_collect___lambda__2___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_CustomEliminators_map___default___lambda__1___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__3; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__7; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__1; static lean_object* l_Lean_Meta_getElimExprInfo___closed__4; lean_object* lean_mk_array(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__2___closed__2; +static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____closed__2; uint8_t lean_usize_dec_eq(size_t, size_t); static lean_object* l_Lean_Meta_instReprCustomEliminator___closed__1; lean_object* l_Lean_throwError___at___private_Lean_Meta_Basic_0__Lean_Meta_processPostponedStep___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_getElimExprInfo___lambda__4___closed__3; uint8_t l_Lean_Name_isAnonymous(lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__7___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__11; +LEAN_EXPORT lean_object* l_Lean_Meta_getCustomEliminator_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____spec__1(lean_object*); static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__13; lean_object* l_Lean_Expr_bvar___override(lean_object*); static lean_object* l_Lean_Meta_instReprCustomEliminators___closed__1; lean_object* lean_array_fget(lean_object*, lean_object*); +static lean_object* l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8___closed__2; +static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____closed__4; lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__9(lean_object*); -LEAN_EXPORT lean_object* l_Array_isEqvAux___at_Lean_Meta_getCustomEliminator_x3f___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__5___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8___closed__4; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____closed__3; +LEAN_EXPORT lean_object* l_Array_isEqvAux___at_Lean_Meta_getCustomEliminator_x3f___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__17; lean_object* l_Lean_Expr_fvarId_x21(lean_object*); +static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___closed__5; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__6___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_CustomEliminators_map___default___closed__1; LEAN_EXPORT lean_object* l_Lean_AssocList_replace___at_Lean_Meta_addCustomEliminatorEntry___spec__16(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__1; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____closed__2; +static lean_object* l_Lean_Meta_CustomEliminators_map___default___closed__9; static lean_object* l_Lean_Meta_getElimExprInfo___lambda__4___closed__1; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_addImplicitTargets___spec__2___closed__5; LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Meta_addCustomEliminatorEntry___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__12; +static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__3; +static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___closed__4; +static lean_object* l_Lean_Meta_CustomEliminators_map___default___closed__10; static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__28; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39_(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__2___closed__1; lean_object* l_Lean_stringToMessageData(lean_object*); +static lean_object* l_Lean_SMap_toList___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__1___closed__1; +static lean_object* l_Lean_Meta_CustomEliminators_map___default___closed__8; LEAN_EXPORT lean_object* l_Lean_exprDependsOn___at_Lean_Meta_mkCustomEliminator___spec__1(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_getElimExprInfo___spec__7___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Meta_mkCustomEliminator___spec__18___closed__3; -LEAN_EXPORT lean_object* l_Lean_Meta_mkCustomEliminator___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_mkCustomEliminator___lambda__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__19; LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Meta_addCustomEliminatorEntry___spec__14___at_Lean_Meta_addCustomEliminatorEntry___spec__15(lean_object*, lean_object*); static lean_object* l_Lean_Meta_addImplicitTargets_collect___lambda__2___closed__3; @@ -129,67 +145,65 @@ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_mkCustomEliminato extern lean_object* l_Lean_Expr_instBEqExpr; static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__20; static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__9; -static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__5; lean_object* l_Lean_ScopedEnvExtension_addEntry___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_instBEqProd___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_addImplicitTargets_collect___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_AssocList_contains___at_Lean_Meta_addCustomEliminatorEntry___spec__10___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__4; LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at_Lean_Meta_mkCustomEliminator___spec__5(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Option_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getElimExprInfo___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_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__7; LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_Meta_mkCustomEliminator___spec__14___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_getElimExprInfo___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Option_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addCustomEliminator___spec__1___boxed(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_mkCustomEliminator___spec__8(lean_object*, lean_object*, size_t, size_t); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__20; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_getCustomEliminator_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_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_mkCustomEliminator___spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_instInhabitedNat; +LEAN_EXPORT lean_object* l_Lean_Meta_mkCustomEliminator___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_altArity(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addCustomEliminator___spec__1___closed__10; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__5(lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_addCustomEliminatorEntry___spec__7(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMapImp_expand___at_Lean_Meta_addCustomEliminatorEntry___spec__12(lean_object*, lean_object*); -static lean_object* l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8___closed__1; -static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__3; uint8_t l_Lean_Expr_hasMVar(lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_getElimExprInfo___spec__5(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Meta_getElimExprInfo___lambda__2(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_Meta_addImplicitTargets___spec__2___closed__1; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__6; -LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__3(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__4(lean_object*); static lean_object* l_Lean_Meta_getElimExprInfo___closed__3; lean_object* l_instHashableArray___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Meta_mkCustomEliminator___spec__18___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__6; LEAN_EXPORT lean_object* l_Option_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_mkCustomEliminator___spec__15(lean_object*, lean_object*, size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_SMap_switch___at_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getElimExprInfo___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__7; LEAN_EXPORT lean_object* l_Array_isEqvAux___at_Lean_Meta_addCustomEliminatorEntry___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__4___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_getCustomEliminator_x3f___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__12(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_Meta_mkCustomEliminator___spec__3___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__10; LEAN_EXPORT lean_object* l_Lean_Meta_getElimExprInfo___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_Meta_getElimExprInfo___closed__6; static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__8; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_getElimExprInfo___spec__3(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_throwError___at_Lean_Meta_getElimExprInfo___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_instReprElimInfo___closed__1; -LEAN_EXPORT lean_object* l_Lean_Meta_addCustomEliminator(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_addCustomEliminator(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ScopedEnvExtension_addScopedEntry___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_getElimExprInfo___lambda__4___closed__2; uint8_t l_Lean_Expr_isSort(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__4; +LEAN_EXPORT uint8_t l_Lean_Meta_CustomEliminators_map___default___lambda__1(uint8_t, uint8_t); static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__9; lean_object* lean_st_ref_take(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__6; -static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__9; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____closed__1; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Meta_getElimExprInfo___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_addImplicitTargets___spec__2___closed__3; @@ -197,6 +211,7 @@ static lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_getElimExprInfo___spec uint8_t lean_expr_eqv(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__22; LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Meta_getCustomEliminator_x3f___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__9; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_mkCustomEliminator___spec__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Option_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____spec__1___closed__3; lean_object* l_Lean_registerSimpleScopedEnvExtension___rarg(lean_object*, lean_object*); @@ -210,141 +225,143 @@ LEAN_EXPORT lean_object* l_Lean_SMap_find_x3f___at_Lean_Meta_getCustomEliminator LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Meta_mkCustomEliminator___spec__18(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Meta_CustomEliminators_map___default___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_addImplicitTargets_collect___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__5; -LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__12(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addCustomEliminator___spec__1___closed__4; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__3; +LEAN_EXPORT lean_object* l_Lean_SMap_fold___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__2(lean_object*); +static size_t l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__2; static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__12; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991_(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isMVar(lean_object*); +lean_object* l_instHashableBool___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getCustomEliminators(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_SMap_toList___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__1___lambda__1(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__2; static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__15; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__3; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237_(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_getCustomEliminator_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_getCustomEliminator_x3f(lean_object*, uint8_t, 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*); static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addCustomEliminator___spec__1___closed__6; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__14; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Meta_mkCustomEliminator___spec__19___lambda__1(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_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__6___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_instInhabitedCustomEliminators___closed__2; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__12; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____closed__4; static lean_object* l_Lean_Meta_addImplicitTargets_collect___closed__1; static lean_object* l_Lean_Meta_instReprElimAltInfo___closed__1; lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l_Lean_Meta_getCustomEliminator_x3f___closed__1; LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_Meta_getCustomEliminator_x3f___spec__9(lean_object*, lean_object*); +static size_t l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__4; lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_MetavarContext_getExprAssignmentDomain___spec__2___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_getElimExprInfo___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_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_Meta_mkCustomEliminator___spec__10(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__11; +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getCustomEliminators___boxed(lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148_(lean_object*, lean_object*); static lean_object* l_Lean_Meta_addCustomEliminator___closed__1; lean_object* lean_array_to_list(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__5; 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; lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__12; -static lean_object* l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8___closed__3; +static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___closed__1; lean_object* l_Lean_Name_append(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__10(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__21; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_addImplicitTargets___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_AssocList_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__3___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_addCustomEliminatorEntry___spec__5___boxed(lean_object*, lean_object*, 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_2237____spec__10___closed__1; static lean_object* l_Lean_Meta_CustomEliminators_map___default___closed__4; -LEAN_EXPORT lean_object* l_Lean_SMap_fold___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__2___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__10; extern lean_object* l_Lean_Expr_instHashableExpr; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__6(lean_object*); extern lean_object* l_Lean_levelZero; -LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__11(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____closed__6; static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__2; static lean_object* l_Option_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____spec__1___closed__4; extern lean_object* l_Lean_instInhabitedExpr; +static lean_object* l_Lean_Meta_CustomEliminators_map___default___closed__7; static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__24; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Meta_mkCustomEliminator___spec__19___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_addImplicitTargets(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_whnfD(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__5; static lean_object* l_Lean_Meta_instInhabitedCustomEliminators___closed__1; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__7___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_instInhabitedCustomEliminator___closed__1; -LEAN_EXPORT lean_object* l_Lean_Meta_mkCustomEliminator(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__11; +LEAN_EXPORT lean_object* l_Lean_Meta_mkCustomEliminator(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_contains___at_Lean_Meta_addImplicitTargets_collect___spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addCustomEliminator___spec__1___closed__8; static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__2; static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__12; -static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__8; +static lean_object* l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8___closed__5; static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__4; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__19; lean_object* l_Lean_instHashableLocalInstance___boxed(lean_object*); static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__17; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____closed__5; lean_object* l_instHashableProd___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMapImp_moveEntries___at_Lean_Meta_addCustomEliminatorEntry___spec__13(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_addImplicitTargets_collect___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_indexOfAux___at_Lean_Meta_getElimExprInfo___spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addCustomEliminator___spec__1___closed__9; uint8_t lean_name_eq(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__9; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_addCustomEliminatorEntry___spec__2(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____closed__7; static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__15; -LEAN_EXPORT lean_object* l_Lean_Meta_getCustomEliminator_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_getCustomEliminator_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____closed__1; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__8; static lean_object* l_Std_Range_forIn_loop___at_Lean_Meta_mkCustomEliminator___spec__19___closed__3; -static lean_object* l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8___closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_instReprElimAltInfo; -static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__7; -LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Meta_addCustomEliminatorEntry___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_SMap_toList___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__1(lean_object*); +LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Meta_addCustomEliminatorEntry___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__23; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_getCustomEliminator_x3f___spec__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_instReprElimInfo; 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_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getCustomEliminators___rarg(lean_object*, lean_object*); -static size_t l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__1; +static uint64_t l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__1; static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addCustomEliminator___spec__1___closed__1; lean_object* l_Lean_LocalDecl_fvarId(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_getCustomEliminator_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_SMap_toList___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_getCustomEliminator_x3f___lambda__1(uint8_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_Tactic_ElimInfo___hyg_2991____lambda__1___closed__4; static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__6; -static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__13; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__7; +static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___closed__2; +LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__13(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_instInhabitedElimInfo; -LEAN_EXPORT lean_object* l_Lean_SMap_toList___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__1(lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____closed__5; LEAN_EXPORT lean_object* l_Array_indexOfAux___at_Lean_Meta_getElimExprInfo___spec__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_instInhabitedElimInfo___closed__1; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__7; LEAN_EXPORT lean_object* l_Array_isEqvAux___at_Lean_Meta_addCustomEliminatorEntry___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_addImplicitTargets_collect___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_instInhabitedCustomEliminator; -static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_instInhabitedElimAltInfo; uint8_t l_Lean_PersistentHashMap_contains___at_Lean_MVarId_isAssigned___spec__1(lean_object*, lean_object*); -LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Meta_getCustomEliminator_x3f___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Meta_getCustomEliminator_x3f___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_ElimInfo_targetsPos___default___closed__1; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Meta_getElimExprInfo___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*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_mkCustomEliminator___spec__7(lean_object*, lean_object*, size_t, size_t); lean_object* lean_usize_to_nat(size_t); size_t lean_hashmap_mk_idx(lean_object*, uint64_t); LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addCustomEliminator___spec__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_SMap_fold___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__2___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__8; static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__25; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_customEliminatorExt; uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__2; lean_object* l_Lean_throwError___at_Lean_LMVarId_getLevel___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofExpr(lean_object*); static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__10; static lean_object* l_Std_Range_forIn_loop___at_Lean_Meta_mkCustomEliminator___spec__19___closed__4; -LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__5; -LEAN_EXPORT lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10(lean_object*, lean_object*); static lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3___closed__3; +static uint64_t l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__3; LEAN_EXPORT lean_object* l_Array_isEqvAux___at_Lean_Meta_getCustomEliminator_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_Tactic_ElimInfo___hyg_2946____closed__4; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__9; -LEAN_EXPORT lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____spec__3(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__6; +static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____closed__5; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__4; lean_object* l_Lean_LocalDecl_userName(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_getElimExprInfo___spec__3___closed__4; @@ -353,84 +370,95 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_getElimExprInfo__ static lean_object* l_Std_Range_forIn_loop___at_Lean_Meta_mkCustomEliminator___spec__18___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_ElimInfo_altsInfo___default; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_mkCustomEliminator___spec__17(lean_object*, lean_object*, size_t, size_t); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__22; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__5___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_contains___at_Lean_Meta_getElimExprInfo___spec__4___boxed(lean_object*, lean_object*); -static lean_object* l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8___closed__4; +LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____spec__3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_getCustomEliminator_x3f___spec__1(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_instInhabitedElimAltInfo___closed__1; static lean_object* l_Lean_Meta_getElimExprInfo___closed__2; uint8_t l_Lean_Name_hasMacroScopes(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_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____closed__7; LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____spec__6(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__4; static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__16; static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__14; +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_isEqvAux___at_Lean_Meta_addCustomEliminatorEntry___spec__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____closed__6; static lean_object* l_Lean_Meta_getElimExprInfo___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_getCustomEliminators___rarg___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__6; static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__19; +static lean_object* l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8___closed__3; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_Meta_mkCustomEliminator___spec__4(lean_object*, lean_object*, lean_object*); lean_object* lean_string_length(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getElimExprInfo___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 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___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__29; +static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__8; +LEAN_EXPORT lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11(lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____closed__3; +LEAN_EXPORT lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___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___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__16; 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_Lean_ScopedEnvExtension_add___at_Lean_Meta_addCustomEliminator___spec__1___closed__15; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); +static lean_object* l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8___closed__1; static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__23; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__6; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__15; lean_object* l_Lean_indentExpr(lean_object*); uint8_t l_Lean_BinderInfo_isExplicit(uint8_t); -static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__2; static size_t l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3___closed__1; lean_object* l_Lean_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__21; +LEAN_EXPORT lean_object* l_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____spec__1(lean_object*); LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____spec__3(lean_object*, lean_object*); lean_object* l_Lean_Meta_mkHasTypeButIsExpectedMsg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__31; lean_object* l_Std_Format_joinSep___at_Prod_repr___spec__1(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at_Lean_Meta_mkCustomEliminator___spec__13(lean_object*, lean_object*); uint64_t l_Lean_Name_hash___override(lean_object*); +LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____spec__2(lean_object*, lean_object*); static lean_object* l_Lean_Meta_instInhabitedElimInfo___closed__2; 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_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__10; lean_object* l_Lean_LocalDecl_type(lean_object*); lean_object* l_Repr_addAppParen(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_Meta_mkCustomEliminator___spec__11(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__10; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__4___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__13; LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Meta_CustomEliminators_map___default___spec__1___boxed(lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____boxed(lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); lean_object* l_Lean_ScopedEnvExtension_getState___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_getElimExprInfo___closed__1; lean_object* l_Lean_Expr_getAppFn(lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__7(lean_object*); LEAN_EXPORT uint8_t l_Lean_AssocList_contains___at_Lean_Meta_addCustomEliminatorEntry___spec__10(lean_object*, lean_object*); lean_object* lean_nat_mul(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__7; LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Meta_addCustomEliminatorEntry___spec__14(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298_(lean_object*); lean_object* l_Lean_Meta_InfoCacheKey_instHashableInfoCacheKey___boxed(lean_object*); static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_altArity___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addCustomEliminator___spec__1___closed__17; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__2; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____closed__4; static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__18; lean_object* l_Lean_PersistentHashMap_mkCollisionNode___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static size_t l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3___closed__2; -LEAN_EXPORT lean_object* l_Array_isEqvAux___at_Lean_Meta_addCustomEliminatorEntry___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_isEqvAux___at_Lean_Meta_addCustomEliminatorEntry___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); uint8_t l_Lean_LocalDecl_binderInfo(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__6___rarg(lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* l_Lean_ScopedEnvExtension_addLocalEntry___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__4; -static lean_object* l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8___closed__6; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__15; -LEAN_EXPORT lean_object* l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8(lean_object*, lean_object*); lean_object* l_Lean_registerBuiltinAttribute(lean_object*, lean_object*); static lean_object* l_Lean_Meta_getElimExprInfo___lambda__1___closed__2; -static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__4; +static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____closed__6; lean_object* l_List_reverse___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_exprDependsOn___at_Lean_Meta_mkCustomEliminator___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_addImplicitTargets___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -438,18 +466,17 @@ static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addCustomElimin static lean_object* l_Lean_Meta_CustomEliminators_map___default___closed__5; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_mkCustomEliminator___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_sub(size_t, size_t); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__7; lean_object* l_Lean_mkHashMapImp___rarg(lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_getElimExprInfo___spec__7(lean_object*, size_t, size_t); static lean_object* l_Lean_Meta_addImplicitTargets_collect___closed__4; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_Meta_mkCustomEliminator___spec__11___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_Meta_mkCustomEliminator___spec__6___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_SMap_toList___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__1___lambda__1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__11; static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__10; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_getCustomEliminator_x3f___spec__1___closed__1; static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__30; -LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__4; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__6(lean_object*); size_t lean_usize_add(size_t, size_t); uint8_t l_Lean_Expr_hasFVar(lean_object*); static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__6; @@ -460,58 +487,65 @@ static lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_getElimExprInfo___spec static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__27; static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addCustomEliminator___spec__1___closed__5; static lean_object* l_Std_Range_forIn_loop___at_Lean_Meta_mkCustomEliminator___spec__19___closed__1; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__2; static lean_object* l_Lean_Meta_addImplicitTargets_collect___closed__3; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_addImplicitTargets___spec__2___closed__4; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__14; LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_Meta_getCustomEliminator_x3f___spec__9___boxed(lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Meta_mkCustomEliminator___spec__18___closed__1; +LEAN_EXPORT lean_object* l_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__9(lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__9; static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addCustomEliminator___spec__1___closed__12; 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*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__1; LEAN_EXPORT lean_object* l_Array_isEqvAux___at_Lean_Meta_getCustomEliminator_x3f___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_shift_left(size_t, size_t); +lean_object* l_instBEq___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3(lean_object*, size_t, size_t, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__3; LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_Meta_mkCustomEliminator___spec__13___boxed(lean_object*, lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Meta_addCustomEliminatorEntry___spec__11(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_mkCustomEliminator___spec__4___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__20; static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__1; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__5; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335_(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___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__21; LEAN_EXPORT uint8_t l_Array_contains___at_Lean_Meta_addImplicitTargets_collect___spec__1(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__18; lean_object* l___private_Lean_Expr_0__Lean_reprExpr____x40_Lean_Expr___hyg_2961_(lean_object*, lean_object*); LEAN_EXPORT uint64_t l_Array_foldlMUnsafe_fold___at_Lean_Meta_addCustomEliminatorEntry___spec__5(lean_object*, size_t, size_t, uint64_t); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__16; lean_object* l_Lean_Expr_headBeta(lean_object*); lean_object* lean_array_get_size(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__3; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__7; +static lean_object* l_Lean_Meta_CustomEliminators_map___default___closed__6; static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__20; static lean_object* l_Lean_Meta_CustomEliminators_map___default___closed__3; LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at_Lean_Meta_addImplicitTargets___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Meta_addCustomEliminatorEntry___spec__17(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Meta_mkCustomEliminator___spec__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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__6___rarg(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_getCustomEliminator_x3f___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_mkCustomEliminator___spec__9___boxed(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_Lean_SMap_switch___at_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____spec__1(lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_Meta_mkCustomEliminator___spec__3(lean_object*, lean_object*, lean_object*); -static lean_object* l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8___closed__2; -static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__2; +static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____closed__1; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__18; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_getElimExprInfo___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_ElimInfo_targetsPos___default; +static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____closed__3; uint8_t lean_usize_dec_lt(size_t, size_t); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__17; +LEAN_EXPORT lean_object* l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__3(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_Meta_mkCustomEliminator___spec__2___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__5(lean_object*); lean_object* l_Lean_getExprMVarAssignment_x3f___at___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___spec__1(lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_Meta_mkCustomEliminator___spec__2(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at_Lean_Meta_mkCustomEliminator___spec__14(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__10___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__17; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_getElimExprInfo___spec__3___closed__1; static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addCustomEliminator___spec__1___closed__14; @@ -520,42 +554,29 @@ lean_object* l_Array_instBEqArray___rarg___boxed(lean_object*, lean_object*, lea LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_Meta_mkCustomEliminator___spec__5___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__1; uint8_t l_Lean_Expr_isFVar(lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__8; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__21; lean_object* l_Lean_Meta_mkConstWithFreshMVarLevels(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_mkCustomEliminator___spec__16(lean_object*, lean_object*, size_t, size_t); lean_object* l_Lean_Expr_mvarId_x21(lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__2___closed__2; static lean_object* l_Lean_Meta_CustomEliminators_map___default___closed__2; -static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__6; LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____spec__4(lean_object*, lean_object*, lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); lean_object* l_Lean_MessageData_ofName(lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__6; lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); lean_object* lean_expr_instantiate1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__4(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__7; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__8; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__2; -LEAN_EXPORT lean_object* l_Lean_Meta_mkCustomEliminator___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_mkCustomEliminator___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__14; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__5; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_addImplicitTargets_collect(lean_object*, lean_object*, lean_object*, 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_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__7(lean_object*); lean_object* l_Nat_repr(lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__11; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__9; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_Meta_mkCustomEliminator___spec__12___boxed(lean_object*, lean_object*, lean_object*); size_t lean_usize_land(size_t, size_t); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__19; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_getElimExprInfo___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addCustomEliminator___spec__1___closed__3; extern lean_object* l_Lean_instHashableName; lean_object* l___private_Lean_Data_HashMap_0__Lean_numBucketsForCapacity(lean_object*); lean_object* l_Lean_throwError___at___private_Lean_Meta_InferType_0__Lean_Meta_inferProjType___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__5; -static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__6; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_addImplicitTargets_collect___spec__2(lean_object*, lean_object*, size_t, size_t); static lean_object* _init_l_Option_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____spec__1___closed__1() { _start: @@ -4471,13 +4492,15 @@ return x_11; static lean_object* _init_l_Lean_Meta_instInhabitedCustomEliminator___closed__1() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_ElimInfo_targetsPos___default___closed__1; -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); -return x_3; +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l_Lean_Meta_ElimInfo_targetsPos___default___closed__1; +x_3 = lean_box(0); +x_4 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_4, 0, x_2); +lean_ctor_set(x_4, 1, x_3); +lean_ctor_set_uint8(x_4, sizeof(void*)*2, x_1); +return x_4; } } static lean_object* _init_l_Lean_Meta_instInhabitedCustomEliminator() { @@ -4488,7 +4511,7 @@ x_1 = l_Lean_Meta_instInhabitedCustomEliminator___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____spec__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____spec__1(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; @@ -4497,7 +4520,7 @@ x_3 = l_Lean_Name_reprPrec(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_3) == 0) @@ -4528,7 +4551,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____spec__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____spec__2(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -4562,47 +4585,39 @@ lean_inc(x_8); lean_dec(x_1); x_9 = lean_unsigned_to_nat(0u); x_10 = l_Lean_Name_reprPrec(x_8, x_9); -x_11 = l_List_foldl___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____spec__3(x_2, x_10, x_4); +x_11 = l_List_foldl___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____spec__3(x_2, x_10, x_4); return x_11; } } } } -static lean_object* _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("typeNames", 9); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__2() { +static lean_object* _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__1() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__1; +x_1 = l_Lean_Meta_getElimExprInfo___closed__2; 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_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__3() { +static lean_object* _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____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_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__2; +x_2 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__1; 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_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__4() { +static lean_object* _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__3; +x_1 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__2; x_2 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__5; x_3 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -4610,231 +4625,246 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__5() { +static lean_object* _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__4() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("elimName", 8); +x_1 = lean_mk_string_from_bytes("typeNames", 9); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__6() { +static lean_object* _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__5; +x_1 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__4; 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_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__7() { +static lean_object* _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__6() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__12; -x_2 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____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; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("elimName", 8); +return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__8() { +static lean_object* _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__7() { _start: { -lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____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; +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____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_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__9() { +static lean_object* _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__4; -x_2 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__8; -x_3 = lean_alloc_ctor(5, 2, 0); +x_1 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__12; +x_2 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____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_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__10() { +static lean_object* _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__9() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__9; -x_2 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__9; -x_3 = lean_alloc_ctor(5, 2, 0); +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__8; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__11() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157_(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__10; -x_2 = lean_box(1); -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_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__12() { -_start: +uint8_t x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_object* x_13; lean_object* x_14; +x_3 = lean_ctor_get_uint8(x_1, sizeof(void*)*2); +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +x_5 = lean_array_get_size(x_4); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_nat_dec_eq(x_5, x_6); +lean_dec(x_5); +x_8 = lean_ctor_get(x_1, 1); +lean_inc(x_8); +lean_dec(x_1); +x_9 = l_Lean_Name_reprPrec(x_8, x_6); +x_10 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__5; +x_11 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_9); +x_12 = 0; +x_13 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_13, 0, x_11); +lean_ctor_set_uint8(x_13, sizeof(void*)*1, x_12); +if (x_3 == 0) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__11; -x_2 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__6; -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; -} +lean_object* x_71; +x_71 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__25; +x_14 = x_71; +goto block_70; } -static lean_object* _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__13() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__12; -x_2 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__5; -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; -} +lean_object* x_72; +x_72 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__29; +x_14 = x_72; +goto block_70; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148_(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; -x_3 = lean_ctor_get(x_1, 0); -lean_inc(x_3); -x_4 = lean_array_get_size(x_3); -x_5 = lean_unsigned_to_nat(0u); -x_6 = lean_nat_dec_eq(x_4, x_5); -lean_dec(x_4); -x_7 = lean_ctor_get(x_1, 1); -lean_inc(x_7); -lean_dec(x_1); -x_8 = l_Lean_Name_reprPrec(x_7, x_5); -x_9 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__5; -x_10 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_10, 0, x_9); -lean_ctor_set(x_10, 1, x_8); -x_11 = 0; -x_12 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_12, 0, x_10); -lean_ctor_set_uint8(x_12, sizeof(void*)*1, x_11); -if (x_6 == 0) +block_70: { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; 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_13 = lean_array_to_list(lean_box(0), x_3); -x_14 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__15; -x_15 = l_Std_Format_joinSep___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____spec__2(x_13, x_14); -x_16 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__19; -x_17 = lean_alloc_ctor(5, 2, 0); +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_15 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__12; +x_16 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_14); +x_17 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_17, 0, x_16); -lean_ctor_set(x_17, 1, x_15); -x_18 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__21; +lean_ctor_set_uint8(x_17, sizeof(void*)*1, x_12); +x_18 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__3; x_19 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_19, 0, x_17); -lean_ctor_set(x_19, 1, x_18); -x_20 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__18; -x_21 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_19); -x_22 = 1; -x_23 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_17); +x_20 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__9; +x_21 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +x_22 = lean_box(1); +x_23 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_23, 0, x_21); -lean_ctor_set_uint8(x_23, sizeof(void*)*1, x_22); -x_24 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__12; -x_25 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_23); -x_26 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set_uint8(x_26, sizeof(void*)*1, x_11); -x_27 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__4; -x_28 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_28, 0, x_27); -lean_ctor_set(x_28, 1, x_26); -x_29 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__9; -x_30 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_30, 0, x_28); -lean_ctor_set(x_30, 1, x_29); -x_31 = lean_box(1); +lean_ctor_set(x_23, 1, x_22); +x_24 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__5; +x_25 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +x_26 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__5; +x_27 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +if (x_7 == 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; uint8_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_28 = lean_array_to_list(lean_box(0), x_4); +x_29 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__15; +x_30 = l_Std_Format_joinSep___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____spec__2(x_28, x_29); +x_31 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__19; x_32 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -x_33 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__6; +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_30); +x_33 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__21; x_34 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_34, 0, x_32); lean_ctor_set(x_34, 1, x_33); -x_35 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__5; -x_36 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -x_37 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_37, 0, x_36); -lean_ctor_set(x_37, 1, x_12); -x_38 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__21; -x_39 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_37); -x_40 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__23; +x_35 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__18; +x_36 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_34); +x_37 = 1; +x_38 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set_uint8(x_38, sizeof(void*)*1, x_37); +x_39 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_39, 0, x_15); +lean_ctor_set(x_39, 1, x_38); +x_40 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set_uint8(x_40, sizeof(void*)*1, x_12); x_41 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 0, x_27); lean_ctor_set(x_41, 1, x_40); -x_42 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__20; -x_43 = lean_alloc_ctor(4, 2, 0); +x_42 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_20); +x_43 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_43, 0, x_42); -lean_ctor_set(x_43, 1, x_41); -x_44 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_44, 0, x_43); -lean_ctor_set_uint8(x_44, sizeof(void*)*1, x_11); -return x_44; -} -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; lean_object* x_51; lean_object* x_52; lean_object* x_53; -lean_dec(x_3); -x_45 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__13; +lean_ctor_set(x_43, 1, x_22); +x_44 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__7; +x_45 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); x_46 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_46, 0, x_45); -lean_ctor_set(x_46, 1, x_12); -x_47 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__21; -x_48 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_48, 0, x_47); -lean_ctor_set(x_48, 1, x_46); -x_49 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__23; -x_50 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_50, 0, x_48); -lean_ctor_set(x_50, 1, x_49); -x_51 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__20; -x_52 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_52, 0, x_51); -lean_ctor_set(x_52, 1, x_50); -x_53 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_46, 1, x_26); +x_47 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_13); +x_48 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__21; +x_49 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_47); +x_50 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__23; +x_51 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_50); +x_52 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__20; +x_53 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_53, 0, x_52); -lean_ctor_set_uint8(x_53, sizeof(void*)*1, x_11); -return x_53; +lean_ctor_set(x_53, 1, x_51); +x_54 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_54, 0, x_53); +lean_ctor_set_uint8(x_54, sizeof(void*)*1, x_12); +return x_54; +} +else +{ +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_dec(x_4); +x_55 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__9; +x_56 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_56, 0, x_27); +lean_ctor_set(x_56, 1, x_55); +x_57 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_20); +x_58 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_58, 0, x_57); +lean_ctor_set(x_58, 1, x_22); +x_59 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__7; +x_60 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_60, 0, x_58); +lean_ctor_set(x_60, 1, x_59); +x_61 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_61, 0, x_60); +lean_ctor_set(x_61, 1, x_26); +x_62 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_62, 0, x_61); +lean_ctor_set(x_62, 1, x_13); +x_63 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__21; +x_64 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_64, 0, x_63); +lean_ctor_set(x_64, 1, x_62); +x_65 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__23; +x_66 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_66, 0, x_64); +lean_ctor_set(x_66, 1, x_65); +x_67 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__20; +x_68 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_68, 0, x_67); +lean_ctor_set(x_68, 1, x_66); +x_69 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_69, 0, x_68); +lean_ctor_set_uint8(x_69, sizeof(void*)*1, x_12); +return x_69; +} } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148_(x_1, x_2); +x_3 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157_(x_1, x_2); lean_dec(x_2); return x_3; } @@ -4843,7 +4873,7 @@ static lean_object* _init_l_Lean_Meta_instReprCustomEliminator___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____boxed), 2, 0); return x_1; } } @@ -4863,22 +4893,44 @@ x_2 = l_Lean_mkHashMapImp___rarg(x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta_CustomEliminators_map___default___closed__1() { +LEAN_EXPORT uint8_t l_Lean_Meta_CustomEliminators_map___default___lambda__1(uint8_t x_1, uint8_t x_2) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Name_instBEqName; -x_2 = lean_alloc_closure((void*)(l_Array_instBEqArray___rarg___boxed), 3, 1); -lean_closure_set(x_2, 0, x_1); +if (x_1 == 0) +{ +if (x_2 == 0) +{ +uint8_t x_3; +x_3 = 1; +return x_3; +} +else +{ +uint8_t x_4; +x_4 = 0; +return x_4; +} +} +else +{ return x_2; } } +} +static lean_object* _init_l_Lean_Meta_CustomEliminators_map___default___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_CustomEliminators_map___default___lambda__1___boxed), 2, 0); +return x_1; +} +} static lean_object* _init_l_Lean_Meta_CustomEliminators_map___default___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_instHashableName; -x_2 = lean_alloc_closure((void*)(l_instHashableArray___rarg___boxed), 2, 1); +x_1 = l_Lean_Meta_CustomEliminators_map___default___closed__1; +x_2 = lean_alloc_closure((void*)(l_instBEq___rarg), 3, 1); lean_closure_set(x_2, 0, x_1); return x_2; } @@ -4886,41 +4938,93 @@ return x_2; static lean_object* _init_l_Lean_Meta_CustomEliminators_map___default___closed__3() { _start: { -lean_object* x_1; -x_1 = l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_box(0), lean_box(0)); -return x_1; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Name_instBEqName; +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_Meta_CustomEliminators_map___default___closed__4() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_CustomEliminators_map___default___closed__3; -x_2 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_CustomEliminators_map___default___closed__2; +x_2 = l_Lean_Meta_CustomEliminators_map___default___closed__3; +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_Meta_CustomEliminators_map___default___closed__5() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_CustomEliminators_map___default___closed__4; -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; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_instHashableName; +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_Meta_CustomEliminators_map___default() { +static lean_object* _init_l_Lean_Meta_CustomEliminators_map___default___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_instHashableBool___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_CustomEliminators_map___default___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_CustomEliminators_map___default___closed__6; +x_2 = l_Lean_Meta_CustomEliminators_map___default___closed__5; +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; +} +} +static lean_object* _init_l_Lean_Meta_CustomEliminators_map___default___closed__8() { +_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_CustomEliminators_map___default___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_CustomEliminators_map___default___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_Meta_CustomEliminators_map___default___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_CustomEliminators_map___default___closed__9; +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_CustomEliminators_map___default() { _start: { lean_object* x_1; lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; x_1 = lean_unsigned_to_nat(8u); x_2 = l_Lean_mkHashMapImp___rarg(x_1); x_3 = 1; -x_4 = l_Lean_Meta_CustomEliminators_map___default___closed__5; +x_4 = l_Lean_Meta_CustomEliminators_map___default___closed__10; x_5 = lean_alloc_ctor(0, 2, 1); lean_ctor_set(x_5, 0, x_2); lean_ctor_set(x_5, 1, x_4); @@ -4937,6 +5041,19 @@ lean_dec(x_1); return x_2; } } +LEAN_EXPORT lean_object* l_Lean_Meta_CustomEliminators_map___default___lambda__1___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; +x_3 = lean_unbox(x_1); +lean_dec(x_1); +x_4 = lean_unbox(x_2); +lean_dec(x_2); +x_5 = l_Lean_Meta_CustomEliminators_map___default___lambda__1(x_3, x_4); +x_6 = lean_box(x_5); +return x_6; +} +} static lean_object* _init_l_Lean_Meta_instInhabitedCustomEliminators___closed__1() { _start: { @@ -4952,7 +5069,7 @@ static lean_object* _init_l_Lean_Meta_instInhabitedCustomEliminators___closed__2 uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = 1; x_2 = l_Lean_Meta_instInhabitedCustomEliminators___closed__1; -x_3 = l_Lean_Meta_CustomEliminators_map___default___closed__5; +x_3 = l_Lean_Meta_CustomEliminators_map___default___closed__10; x_4 = lean_alloc_ctor(0, 2, 1); lean_ctor_set(x_4, 0, x_2); lean_ctor_set(x_4, 1, x_3); @@ -4968,7 +5085,7 @@ x_1 = l_Lean_Meta_instInhabitedCustomEliminators___closed__2; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__3___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__3___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_3) == 0) @@ -4994,15 +5111,15 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__3(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__3(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_AssocList_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__3___rarg), 3, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_AssocList_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__3___rarg), 3, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__4___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__4___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; @@ -5013,15 +5130,15 @@ x_5 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_MetavarContext_getExprAssignm return x_5; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__4(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__4(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__4___rarg), 3, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__4___rarg), 3, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__5___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__5___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; @@ -5032,15 +5149,15 @@ x_5 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_MetavarContext_getExprAssignm return x_5; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__5(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__5(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__5___rarg), 3, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__5___rarg), 3, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__6___rarg(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__6___rarg(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; @@ -5050,7 +5167,7 @@ if (x_6 == 0) lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; x_7 = lean_array_uget(x_2, x_3); lean_inc(x_1); -x_8 = l_Lean_AssocList_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__3___rarg(x_1, x_5, x_7); +x_8 = l_Lean_AssocList_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__3___rarg(x_1, x_5, x_7); x_9 = 1; x_10 = lean_usize_add(x_3, x_9); x_3 = x_10; @@ -5064,15 +5181,15 @@ return x_5; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__6(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__6(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__6___rarg___boxed), 5, 0); +x_2 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__6___rarg___boxed), 5, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__7___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__7___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; @@ -5083,15 +5200,15 @@ x_5 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_MetavarContext_getExprAssignm return x_5; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__7(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__7(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__7___rarg), 3, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__7___rarg), 3, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_SMap_fold___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__2___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_SMap_fold___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__2___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; @@ -5111,7 +5228,7 @@ if (x_9 == 0) lean_object* x_10; lean_dec(x_7); lean_dec(x_6); -x_10 = l_Lean_PersistentHashMap_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__4___rarg(x_4, x_1, x_2); +x_10 = l_Lean_PersistentHashMap_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__4___rarg(x_4, x_1, x_2); return x_10; } else @@ -5123,7 +5240,7 @@ if (x_11 == 0) lean_object* x_12; lean_dec(x_7); lean_dec(x_6); -x_12 = l_Lean_PersistentHashMap_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__5___rarg(x_4, x_1, x_2); +x_12 = l_Lean_PersistentHashMap_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__5___rarg(x_4, x_1, x_2); return x_12; } else @@ -5133,23 +5250,23 @@ x_13 = 0; x_14 = lean_usize_of_nat(x_7); lean_dec(x_7); lean_inc(x_1); -x_15 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__6___rarg(x_1, x_6, x_13, x_14, x_2); +x_15 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__6___rarg(x_1, x_6, x_13, x_14, x_2); lean_dec(x_6); -x_16 = l_Lean_PersistentHashMap_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__7___rarg(x_4, x_1, x_15); +x_16 = l_Lean_PersistentHashMap_foldlM___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__7___rarg(x_4, x_1, x_15); return x_16; } } } } -LEAN_EXPORT lean_object* l_Lean_SMap_fold___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_SMap_fold___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__2(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_SMap_fold___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__2___rarg), 3, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_SMap_fold___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__2___rarg), 3, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_SMap_toList___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_SMap_toList___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; @@ -5162,25 +5279,25 @@ lean_ctor_set(x_5, 1, x_1); return x_5; } } -static lean_object* _init_l_Lean_SMap_toList___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__1___closed__1() { +static lean_object* _init_l_Lean_SMap_toList___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__1___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_SMap_toList___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__1___lambda__1), 3, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_SMap_toList___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__1___lambda__1), 3, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_SMap_toList___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_SMap_toList___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__1(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; x_2 = lean_box(0); -x_3 = l_Lean_SMap_toList___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__1___closed__1; -x_4 = l_Lean_SMap_fold___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__2___rarg(x_3, x_2, x_1); +x_3 = l_Lean_SMap_toList___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__1___closed__1; +x_4 = l_Lean_SMap_fold___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__2___rarg(x_3, x_2, x_1); return x_4; } } -static lean_object* _init_l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__1() { +static lean_object* _init_l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___closed__1() { _start: { lean_object* x_1; @@ -5188,35 +5305,35 @@ x_1 = lean_mk_string_from_bytes("(", 1); return x_1; } } -static lean_object* _init_l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__2() { +static lean_object* _init_l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__1; +x_1 = l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___closed__1; x_2 = lean_string_length(x_1); return x_2; } } -static lean_object* _init_l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__3() { +static lean_object* _init_l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__2; +x_1 = l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___closed__2; x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__4() { +static lean_object* _init_l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__1; +x_1 = l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___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_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__5() { +static lean_object* _init_l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___closed__5() { _start: { lean_object* x_1; @@ -5224,82 +5341,88 @@ x_1 = lean_mk_string_from_bytes(")", 1); return x_1; } } -static lean_object* _init_l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__6() { +static lean_object* _init_l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__5; +x_1 = l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___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_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__7() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__23; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11(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; lean_object* x_8; lean_object* x_9; +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; uint8_t x_47; x_3 = lean_ctor_get(x_1, 0); lean_inc(x_3); x_4 = lean_ctor_get(x_1, 1); lean_inc(x_4); lean_dec(x_1); -x_5 = lean_array_get_size(x_3); -x_6 = lean_unsigned_to_nat(0u); -x_7 = lean_nat_dec_eq(x_5, x_6); -lean_dec(x_5); -x_8 = lean_box(0); -x_9 = l_Lean_Name_reprPrec(x_4, x_6); -if (x_7 == 0) +x_5 = lean_box(0); +x_6 = lean_array_get_size(x_4); +x_7 = lean_unsigned_to_nat(0u); +x_8 = lean_nat_dec_eq(x_6, x_7); +lean_dec(x_6); +x_47 = lean_unbox(x_3); +lean_dec(x_3); +if (x_47 == 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; 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; lean_object* x_30; uint8_t x_31; lean_object* x_32; -x_10 = lean_array_to_list(lean_box(0), x_3); -x_11 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__15; -x_12 = l_Std_Format_joinSep___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____spec__2(x_10, x_11); -x_13 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__19; -x_14 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_12); -x_15 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__21; -x_16 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_16, 0, x_14); -lean_ctor_set(x_16, 1, x_15); -x_17 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__18; -x_18 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_16); -x_19 = 1; -x_20 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_20, 0, x_18); -lean_ctor_set_uint8(x_20, sizeof(void*)*1, x_19); -x_21 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_8); +lean_object* x_48; +x_48 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__25; +x_9 = x_48; +goto block_46; +} +else +{ +lean_object* x_49; +x_49 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__29; +x_9 = x_49; +goto block_46; +} +block_46: +{ +lean_object* x_10; +x_10 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_5); +if (x_8 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; 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; uint8_t x_31; lean_object* x_32; +x_11 = lean_array_to_list(lean_box(0), x_4); +x_12 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__15; +x_13 = l_Std_Format_joinSep___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____spec__2(x_11, x_12); +x_14 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__19; +x_15 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_13); +x_16 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__21; +x_17 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set(x_17, 1, x_16); +x_18 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__18; +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 = 1; +x_21 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set_uint8(x_21, sizeof(void*)*1, x_20); x_22 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_22, 0, x_9); -lean_ctor_set(x_22, 1, x_21); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_10); x_23 = l_List_reverse___rarg(x_22); -x_24 = l_Std_Format_joinSep___at_Prod_repr___spec__1(x_23, x_11); -x_25 = l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__4; +x_24 = l_Std_Format_joinSep___at_Prod_repr___spec__1(x_23, x_12); +x_25 = l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___closed__4; x_26 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_26, 0, x_25); lean_ctor_set(x_26, 1, x_24); -x_27 = l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__6; +x_27 = l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___closed__6; x_28 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_28, 0, x_26); lean_ctor_set(x_28, 1, x_27); -x_29 = l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__3; +x_29 = l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___closed__3; x_30 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_30, 0, x_29); lean_ctor_set(x_30, 1, x_28); @@ -5312,23 +5435,23 @@ 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; uint8_t x_44; lean_object* x_45; -lean_dec(x_3); -x_33 = l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__7; +lean_dec(x_4); +x_33 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__23; x_34 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_34, 0, x_9); -lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_10); x_35 = l_List_reverse___rarg(x_34); x_36 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__15; x_37 = l_Std_Format_joinSep___at_Prod_repr___spec__1(x_35, x_36); -x_38 = l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__4; +x_38 = l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___closed__4; x_39 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_39, 0, x_38); lean_ctor_set(x_39, 1, x_37); -x_40 = l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__6; +x_40 = l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___closed__6; x_41 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_41, 0, x_39); lean_ctor_set(x_41, 1, x_40); -x_42 = l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__3; +x_42 = l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___closed__3; x_43 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_43, 0, x_42); lean_ctor_set(x_43, 1, x_41); @@ -5340,16 +5463,58 @@ return x_45; } } } -LEAN_EXPORT lean_object* l_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__9(lean_object* x_1) { +} +LEAN_EXPORT lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____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; 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; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 1); +lean_inc(x_4); +lean_dec(x_1); +x_5 = lean_unsigned_to_nat(0u); +x_6 = l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11(x_3, x_5); +x_7 = lean_box(0); +x_8 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_8, 0, x_6); +lean_ctor_set(x_8, 1, x_7); +x_9 = l_Lean_Name_reprPrec(x_4, x_5); +x_10 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_8); +x_11 = l_List_reverse___rarg(x_10); +x_12 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__15; +x_13 = l_Std_Format_joinSep___at_Prod_repr___spec__1(x_11, x_12); +x_14 = l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___closed__4; +x_15 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_13); +x_16 = l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___closed__6; +x_17 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set(x_17, 1, x_16); +x_18 = l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___closed__3; +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 = 0; +x_21 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set_uint8(x_21, sizeof(void*)*1, x_20); +return x_21; +} +} +LEAN_EXPORT lean_object* l_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__9(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10(x_1, x_2); +x_3 = l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__10(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__13(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_3) == 0) @@ -5370,7 +5535,7 @@ x_6 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_6, 0, x_2); lean_ctor_set(x_6, 1, x_1); x_7 = lean_unsigned_to_nat(0u); -x_8 = l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10(x_4, x_7); +x_8 = l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__10(x_4, x_7); x_9 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_9, 0, x_6); lean_ctor_set(x_9, 1, x_8); @@ -5380,7 +5545,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__11(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__12(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -5403,7 +5568,7 @@ x_5 = lean_ctor_get(x_1, 0); lean_inc(x_5); lean_dec(x_1); x_6 = lean_unsigned_to_nat(0u); -x_7 = l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10(x_5, x_6); +x_7 = l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__10(x_5, x_6); return x_7; } else @@ -5413,14 +5578,14 @@ x_8 = lean_ctor_get(x_1, 0); lean_inc(x_8); lean_dec(x_1); x_9 = lean_unsigned_to_nat(0u); -x_10 = l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10(x_8, x_9); -x_11 = l_List_foldl___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__12(x_2, x_10, x_4); +x_10 = l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__10(x_8, x_9); +x_11 = l_List_foldl___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__13(x_2, x_10, x_4); return x_11; } } } } -static lean_object* _init_l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8___closed__1() { +static lean_object* _init_l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8___closed__1() { _start: { lean_object* x_1; @@ -5428,17 +5593,17 @@ x_1 = lean_mk_string_from_bytes("[]", 2); return x_1; } } -static lean_object* _init_l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8___closed__2() { +static lean_object* _init_l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8___closed__1; +x_1 = l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8___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_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8___closed__3() { +static lean_object* _init_l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8___closed__3() { _start: { lean_object* x_1; @@ -5446,49 +5611,49 @@ x_1 = lean_mk_string_from_bytes("[", 1); return x_1; } } -static lean_object* _init_l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8___closed__4() { +static lean_object* _init_l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8___closed__3; +x_1 = l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8___closed__3; x_2 = lean_string_length(x_1); return x_2; } } -static lean_object* _init_l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8___closed__5() { +static lean_object* _init_l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8___closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8___closed__4; +x_1 = l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8___closed__4; x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8___closed__6() { +static lean_object* _init_l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8___closed__3; +x_1 = l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8___closed__3; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) { lean_object* x_3; -x_3 = l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8___closed__2; +x_3 = l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8___closed__2; return x_3; } else { lean_object* x_4; lean_object* x_5; 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; x_4 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__15; -x_5 = l_Std_Format_joinSep___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__11(x_1, x_4); -x_6 = l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8___closed__6; +x_5 = l_Std_Format_joinSep___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__12(x_1, x_4); +x_6 = l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8___closed__6; x_7 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); @@ -5496,7 +5661,7 @@ x_8 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lea x_9 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_9, 0, x_7); lean_ctor_set(x_9, 1, x_8); -x_10 = l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8___closed__5; +x_10 = l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8___closed__5; x_11 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -5508,7 +5673,7 @@ return x_13; } } } -static lean_object* _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__1() { +static lean_object* _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____closed__1() { _start: { lean_object* x_1; @@ -5516,33 +5681,33 @@ x_1 = lean_mk_string_from_bytes("map", 3); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__2() { +static lean_object* _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__1; +x_1 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____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_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__3() { +static lean_object* _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__2; +x_2 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____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_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__4() { +static lean_object* _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__3; +x_1 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____closed__3; x_2 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__5; x_3 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -5550,7 +5715,7 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__5() { +static lean_object* _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____closed__5() { _start: { lean_object* x_1; lean_object* x_2; @@ -5559,7 +5724,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__6() { +static lean_object* _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____closed__6() { _start: { lean_object* x_1; @@ -5567,30 +5732,30 @@ x_1 = lean_mk_string_from_bytes(".toSMap", 7); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__7() { +static lean_object* _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__6; +x_1 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____closed__6; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270_(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; uint8_t x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_3 = l_Lean_SMap_toList___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__1(x_1); +x_3 = l_Lean_SMap_toList___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__1(x_1); x_4 = lean_unsigned_to_nat(1024u); -x_5 = l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8(x_3, x_4); -x_6 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__7; +x_5 = l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8(x_3, x_4); +x_6 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____closed__7; x_7 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_7, 0, x_5); lean_ctor_set(x_7, 1, x_6); x_8 = lean_unsigned_to_nat(0u); x_9 = l_Repr_addAppParen(x_7, x_8); -x_10 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__5; +x_10 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____closed__5; x_11 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -5598,7 +5763,7 @@ x_12 = 0; x_13 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_13, 0, x_11); lean_ctor_set_uint8(x_13, sizeof(void*)*1, x_12); -x_14 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__4; +x_14 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____closed__4; x_15 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_15, 0, x_14); lean_ctor_set(x_15, 1, x_13); @@ -5620,7 +5785,7 @@ lean_ctor_set_uint8(x_22, sizeof(void*)*1, x_12); return x_22; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__6___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__6___rarg___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; @@ -5628,34 +5793,43 @@ 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___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__6___rarg(x_1, x_2, x_6, x_7, x_5); +x_8 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__6___rarg(x_1, x_2, x_6, x_7, x_5); lean_dec(x_2); return x_8; } } -LEAN_EXPORT lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__10___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10(x_1, x_2); +x_3 = l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__10(x_1, x_2); lean_dec(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8(x_1, x_2); +x_3 = l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8(x_1, x_2); lean_dec(x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237_(x_1, x_2); +x_3 = l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270_(x_1, x_2); lean_dec(x_2); return x_3; } @@ -5664,7 +5838,7 @@ static lean_object* _init_l_Lean_Meta_instReprCustomEliminators___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____boxed), 2, 0); return x_1; } } @@ -5700,11 +5874,40 @@ return x_4; } } } -static size_t _init_l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__1() { +static uint64_t _init_l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__1() { +_start: +{ +uint64_t x_1; uint64_t x_2; uint64_t x_3; +x_1 = 7; +x_2 = 13; +x_3 = lean_uint64_mix_hash(x_2, x_1); +return x_3; +} +} +static size_t _init_l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__2() { _start: { uint64_t x_1; size_t x_2; +x_1 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__1; +x_2 = lean_uint64_to_usize(x_1); +return x_2; +} +} +static uint64_t _init_l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__3() { +_start: +{ +uint64_t x_1; uint64_t x_2; uint64_t x_3; x_1 = 7; +x_2 = 11; +x_3 = lean_uint64_mix_hash(x_2, x_1); +return x_3; +} +} +static size_t _init_l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__4() { +_start: +{ +uint64_t x_1; size_t x_2; +x_1 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__3; x_2 = lean_uint64_to_usize(x_1); return x_2; } @@ -5723,67 +5926,133 @@ return x_6; } else { -lean_object* x_9; lean_object* x_10; uint64_t x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; size_t x_15; size_t x_16; size_t x_17; size_t x_18; lean_object* x_19; lean_object* x_20; +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; uint8_t x_23; x_9 = lean_array_fget(x_2, x_5); x_10 = lean_array_fget(x_3, x_5); -x_11 = 7; -x_12 = lean_array_get_size(x_9); -x_13 = lean_unsigned_to_nat(0u); -x_14 = lean_nat_dec_lt(x_13, x_12); -x_15 = 1; -x_16 = lean_usize_sub(x_1, x_15); -x_17 = 5; -x_18 = lean_usize_mul(x_17, x_16); -x_19 = lean_unsigned_to_nat(1u); -x_20 = lean_nat_add(x_5, x_19); +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); -if (x_14 == 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_18); +x_21 = lean_unsigned_to_nat(0u); +x_22 = lean_nat_dec_lt(x_21, x_20); +x_23 = lean_unbox(x_17); +lean_dec(x_17); +if (x_23 == 0) { -size_t x_21; size_t x_22; lean_object* x_23; -lean_dec(x_12); -x_21 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__1; -x_22 = lean_usize_shift_right(x_21, x_18); -x_23 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3(x_6, x_22, x_1, x_9, x_10); +if (x_22 == 0) +{ +size_t x_24; size_t x_25; lean_object* x_26; +lean_dec(x_20); +lean_dec(x_18); +x_24 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__2; +x_25 = lean_usize_shift_right(x_24, x_14); +x_26 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3(x_6, x_25, x_1, x_9, x_10); x_4 = lean_box(0); -x_5 = x_20; -x_6 = x_23; +x_5 = x_16; +x_6 = x_26; goto _start; } else { -uint8_t x_25; -x_25 = lean_nat_dec_le(x_12, x_12); -if (x_25 == 0) +uint64_t x_28; uint8_t x_29; +x_28 = 13; +x_29 = lean_nat_dec_le(x_20, x_20); +if (x_29 == 0) { -size_t x_26; size_t x_27; lean_object* x_28; -lean_dec(x_12); -x_26 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__1; -x_27 = lean_usize_shift_right(x_26, x_18); -x_28 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3(x_6, x_27, x_1, x_9, x_10); +size_t x_30; size_t x_31; lean_object* x_32; +lean_dec(x_20); +lean_dec(x_18); +x_30 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__2; +x_31 = lean_usize_shift_right(x_30, x_14); +x_32 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3(x_6, x_31, x_1, x_9, x_10); x_4 = lean_box(0); -x_5 = x_20; -x_6 = x_28; +x_5 = x_16; +x_6 = x_32; goto _start; } else { -size_t x_30; size_t x_31; uint64_t x_32; size_t x_33; size_t x_34; lean_object* x_35; -x_30 = 0; -x_31 = lean_usize_of_nat(x_12); -lean_dec(x_12); -x_32 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_addCustomEliminatorEntry___spec__5(x_9, x_30, x_31, x_11); -x_33 = lean_uint64_to_usize(x_32); -x_34 = lean_usize_shift_right(x_33, x_18); -x_35 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3(x_6, x_34, x_1, x_9, x_10); +size_t x_34; size_t x_35; uint64_t x_36; uint64_t x_37; size_t x_38; size_t x_39; lean_object* x_40; +x_34 = 0; +x_35 = lean_usize_of_nat(x_20); +lean_dec(x_20); +x_36 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_addCustomEliminatorEntry___spec__5(x_18, x_34, x_35, x_19); +lean_dec(x_18); +x_37 = lean_uint64_mix_hash(x_28, x_36); +x_38 = lean_uint64_to_usize(x_37); +x_39 = lean_usize_shift_right(x_38, x_14); +x_40 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3(x_6, x_39, x_1, x_9, x_10); +x_4 = lean_box(0); +x_5 = x_16; +x_6 = x_40; +goto _start; +} +} +} +else +{ +if (x_22 == 0) +{ +size_t x_42; size_t x_43; lean_object* x_44; +lean_dec(x_20); +lean_dec(x_18); +x_42 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__4; +x_43 = lean_usize_shift_right(x_42, x_14); +x_44 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3(x_6, x_43, x_1, x_9, x_10); +x_4 = lean_box(0); +x_5 = x_16; +x_6 = x_44; +goto _start; +} +else +{ +uint64_t x_46; uint8_t x_47; +x_46 = 11; +x_47 = lean_nat_dec_le(x_20, x_20); +if (x_47 == 0) +{ +size_t x_48; size_t x_49; lean_object* x_50; +lean_dec(x_20); +lean_dec(x_18); +x_48 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__4; +x_49 = lean_usize_shift_right(x_48, x_14); +x_50 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3(x_6, x_49, x_1, x_9, x_10); +x_4 = lean_box(0); +x_5 = x_16; +x_6 = x_50; +goto _start; +} +else +{ +size_t x_52; size_t x_53; uint64_t x_54; uint64_t x_55; size_t x_56; size_t x_57; lean_object* x_58; +x_52 = 0; +x_53 = lean_usize_of_nat(x_20); +lean_dec(x_20); +x_54 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_addCustomEliminatorEntry___spec__5(x_18, x_52, x_53, x_19); +lean_dec(x_18); +x_55 = lean_uint64_mix_hash(x_46, x_54); +x_56 = lean_uint64_to_usize(x_55); +x_57 = lean_usize_shift_right(x_56, x_14); +x_58 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3(x_6, x_57, x_1, x_9, x_10); x_4 = lean_box(0); -x_5 = x_20; -x_6 = x_35; +x_5 = x_16; +x_6 = x_58; goto _start; } } } } } +} LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Meta_addCustomEliminatorEntry___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: { @@ -5826,49 +6095,48 @@ goto _start; } } } -LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Meta_addCustomEliminatorEntry___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_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Meta_addCustomEliminatorEntry___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) { _start: { -lean_object* x_9; uint8_t x_10; -x_9 = lean_array_get_size(x_6); -x_10 = lean_nat_dec_lt(x_8, x_9); -lean_dec(x_9); -if (x_10 == 0) +lean_object* x_7; uint8_t x_8; +x_7 = lean_array_get_size(x_4); +x_8 = lean_nat_dec_lt(x_6, x_7); +lean_dec(x_7); +if (x_8 == 0) { -uint8_t x_11; -lean_dec(x_8); -x_11 = 1; -return x_11; +uint8_t x_9; +lean_dec(x_6); +x_9 = 1; +return x_9; } else { -lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_12 = lean_array_fget(x_6, x_8); -x_13 = lean_array_fget(x_7, x_8); -x_14 = lean_name_eq(x_12, x_13); -lean_dec(x_13); -lean_dec(x_12); -if (x_14 == 0) +lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_10 = lean_array_fget(x_4, x_6); +x_11 = lean_array_fget(x_5, x_6); +x_12 = lean_name_eq(x_10, x_11); +lean_dec(x_11); +lean_dec(x_10); +if (x_12 == 0) { -uint8_t x_15; -lean_dec(x_8); -x_15 = 0; -return x_15; +uint8_t x_13; +lean_dec(x_6); +x_13 = 0; +return x_13; } else { -lean_object* x_16; lean_object* x_17; -x_16 = lean_unsigned_to_nat(1u); -x_17 = lean_nat_add(x_8, x_16); -lean_dec(x_8); -x_4 = lean_box(0); -x_5 = lean_box(0); -x_8 = x_17; -goto _start; -} -} -} -} +lean_object* x_14; lean_object* x_15; +x_14 = lean_unsigned_to_nat(1u); +x_15 = lean_nat_add(x_6, x_14); +lean_dec(x_6); +x_3 = lean_box(0); +x_6 = x_15; +goto _start; +} +} +} +} LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_addCustomEliminatorEntry___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { @@ -5912,71 +6180,144 @@ return x_16; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +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_46; x_17 = lean_array_fget(x_5, x_2); -x_18 = lean_array_get_size(x_3); -x_19 = lean_array_get_size(x_17); -x_20 = lean_nat_dec_eq(x_18, x_19); -lean_dec(x_19); +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_46 = lean_unbox(x_18); lean_dec(x_18); -if (x_20 == 0) +if (x_46 == 0) { -lean_object* x_21; lean_object* x_22; -lean_dec(x_17); +uint8_t x_47; +x_47 = lean_unbox(x_20); +lean_dec(x_20); +if (x_47 == 0) +{ +uint8_t x_48; +x_48 = 1; +x_22 = x_48; +goto block_45; +} +else +{ +uint8_t x_49; +x_49 = 0; +x_22 = x_49; +goto block_45; +} +} +else +{ +uint8_t x_50; +x_50 = lean_unbox(x_20); +lean_dec(x_20); +if (x_50 == 0) +{ +uint8_t x_51; +x_51 = 0; +x_22 = x_51; +goto block_45; +} +else +{ +uint8_t x_52; +x_52 = 1; +x_22 = x_52; +goto block_45; +} +} +block_45: +{ +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; +lean_dec(x_21); +lean_dec(x_19); lean_dec(x_6); lean_dec(x_5); -x_21 = lean_unsigned_to_nat(1u); -x_22 = lean_nat_add(x_2, x_21); +x_23 = lean_unsigned_to_nat(1u); +x_24 = lean_nat_add(x_2, x_23); lean_dec(x_2); -x_2 = x_22; +x_2 = x_24; goto _start; } else { -lean_object* x_24; uint8_t x_25; -x_24 = lean_unsigned_to_nat(0u); -x_25 = l_Array_isEqvAux___at_Lean_Meta_addCustomEliminatorEntry___spec__8(x_2, x_3, x_5, lean_box(0), lean_box(0), x_3, x_17, x_24); -lean_dec(x_17); -if (x_25 == 0) +lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_26 = lean_array_get_size(x_19); +x_27 = lean_array_get_size(x_21); +x_28 = lean_nat_dec_eq(x_26, x_27); +lean_dec(x_27); +lean_dec(x_26); +if (x_28 == 0) { -lean_object* x_26; lean_object* x_27; +lean_object* x_29; lean_object* x_30; +lean_dec(x_21); +lean_dec(x_19); lean_dec(x_6); lean_dec(x_5); -x_26 = lean_unsigned_to_nat(1u); -x_27 = lean_nat_add(x_2, x_26); +x_29 = lean_unsigned_to_nat(1u); +x_30 = lean_nat_add(x_2, x_29); lean_dec(x_2); -x_2 = x_27; +x_2 = x_30; goto _start; } else { -uint8_t x_29; -x_29 = !lean_is_exclusive(x_1); -if (x_29 == 0) +lean_object* x_32; uint8_t x_33; +x_32 = lean_unsigned_to_nat(0u); +x_33 = l_Array_isEqvAux___at_Lean_Meta_addCustomEliminatorEntry___spec__8(x_19, x_21, lean_box(0), x_19, x_21, x_32); +lean_dec(x_21); +lean_dec(x_19); +if (x_33 == 0) { -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_30 = lean_ctor_get(x_1, 1); -lean_dec(x_30); -x_31 = lean_ctor_get(x_1, 0); -lean_dec(x_31); -x_32 = lean_array_fset(x_5, x_2, x_3); -x_33 = lean_array_fset(x_6, x_2, x_4); +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 +{ +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_33); -lean_ctor_set(x_1, 0, x_32); +lean_ctor_set(x_1, 1, x_41); +lean_ctor_set(x_1, 0, x_40); return x_1; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_dec(x_1); -x_34 = lean_array_fset(x_5, x_2, x_3); -x_35 = lean_array_fset(x_6, x_2, x_4); +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_36 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -return x_36; +x_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; +} +} } } } @@ -6016,543 +6357,597 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCu { 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); +lean_object* x_6; 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_6 = lean_ctor_get(x_1, 0); +lean_inc(x_6); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + x_7 = x_1; +} else { + lean_dec_ref(x_1); + x_7 = lean_box(0); +} x_8 = 1; x_9 = 5; x_10 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3___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_13 = lean_array_get_size(x_6); x_14 = lean_nat_dec_lt(x_12, x_13); lean_dec(x_13); if (x_14 == 0) { +lean_object* x_15; lean_dec(x_12); lean_dec(x_5); lean_dec(x_4); -return x_1; +if (lean_is_scalar(x_7)) { + x_15 = lean_alloc_ctor(0, 1, 0); +} else { + x_15 = x_7; +} +lean_ctor_set(x_15, 0, x_6); +return x_15; } 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)) { +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_array_fget(x_6, x_12); +x_17 = lean_box(0); +x_18 = lean_array_fset(x_6, x_12, x_17); +switch (lean_obj_tag(x_16)) { case 0: { -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; uint8_t x_23; -x_19 = lean_ctor_get(x_15, 0); -x_20 = lean_ctor_get(x_15, 1); -x_21 = lean_array_get_size(x_4); -x_22 = lean_array_get_size(x_19); -x_23 = lean_nat_dec_eq(x_21, x_22); +lean_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; uint8_t x_48; +x_19 = lean_ctor_get(x_16, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_16, 1); +lean_inc(x_20); +if (lean_is_exclusive(x_16)) { + lean_ctor_release(x_16, 0); + lean_ctor_release(x_16, 1); + x_21 = x_16; +} else { + lean_dec_ref(x_16); + x_21 = lean_box(0); +} +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_19, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_19, 1); +lean_inc(x_25); +x_48 = lean_unbox(x_22); lean_dec(x_22); -lean_dec(x_21); -if (x_23 == 0) +if (x_48 == 0) { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -lean_free_object(x_15); -x_24 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_4, x_5); -x_25 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_25, 0, x_24); -x_26 = lean_array_fset(x_17, x_12, x_25); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_26); -return x_1; +uint8_t x_49; +x_49 = lean_unbox(x_24); +lean_dec(x_24); +if (x_49 == 0) +{ +uint8_t x_50; +x_50 = 1; +x_26 = x_50; +goto block_47; } else { -lean_object* x_27; uint8_t x_28; -x_27 = lean_unsigned_to_nat(0u); -x_28 = l_Array_isEqvAux___at_Lean_Meta_addCustomEliminatorEntry___spec__6(x_4, x_19, lean_box(0), x_4, x_19, x_27); -if (x_28 == 0) +uint8_t x_51; +x_51 = 0; +x_26 = x_51; +goto block_47; +} +} +else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; -lean_free_object(x_15); -x_29 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_4, x_5); -x_30 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_30, 0, x_29); -x_31 = lean_array_fset(x_17, x_12, x_30); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_31); -return x_1; +uint8_t x_52; +x_52 = lean_unbox(x_24); +lean_dec(x_24); +if (x_52 == 0) +{ +uint8_t x_53; +x_53 = 0; +x_26 = x_53; +goto block_47; } else { -lean_object* x_32; -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_32 = lean_array_fset(x_17, x_12, x_15); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_32); -return x_1; +uint8_t x_54; +x_54 = 1; +x_26 = x_54; +goto block_47; } } +block_47: +{ +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +lean_dec(x_25); +lean_dec(x_23); +lean_dec(x_21); +x_27 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_4, x_5); +x_28 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_28, 0, x_27); +x_29 = lean_array_fset(x_18, x_12, x_28); +lean_dec(x_12); +if (lean_is_scalar(x_7)) { + x_30 = lean_alloc_ctor(0, 1, 0); +} else { + x_30 = x_7; +} +lean_ctor_set(x_30, 0, x_29); +return x_30; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_33 = lean_ctor_get(x_15, 0); -x_34 = lean_ctor_get(x_15, 1); -lean_inc(x_34); -lean_inc(x_33); -lean_dec(x_15); -x_35 = lean_array_get_size(x_4); -x_36 = lean_array_get_size(x_33); -x_37 = lean_nat_dec_eq(x_35, x_36); -lean_dec(x_36); -lean_dec(x_35); -if (x_37 == 0) +lean_object* x_31; lean_object* x_32; uint8_t x_33; +x_31 = lean_array_get_size(x_23); +x_32 = lean_array_get_size(x_25); +x_33 = lean_nat_dec_eq(x_31, x_32); +lean_dec(x_32); +lean_dec(x_31); +if (x_33 == 0) { -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_33, x_34, x_4, x_5); -x_39 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_39, 0, x_38); -x_40 = lean_array_fset(x_17, x_12, x_39); +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +lean_dec(x_25); +lean_dec(x_23); +lean_dec(x_21); +x_34 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_4, x_5); +x_35 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_35, 0, x_34); +x_36 = lean_array_fset(x_18, x_12, x_35); lean_dec(x_12); -lean_ctor_set(x_1, 0, x_40); -return x_1; +if (lean_is_scalar(x_7)) { + x_37 = lean_alloc_ctor(0, 1, 0); +} else { + x_37 = x_7; +} +lean_ctor_set(x_37, 0, x_36); +return x_37; } else { -lean_object* x_41; uint8_t x_42; -x_41 = lean_unsigned_to_nat(0u); -x_42 = l_Array_isEqvAux___at_Lean_Meta_addCustomEliminatorEntry___spec__6(x_4, x_33, lean_box(0), x_4, x_33, x_41); -if (x_42 == 0) +lean_object* x_38; uint8_t x_39; +x_38 = lean_unsigned_to_nat(0u); +x_39 = l_Array_isEqvAux___at_Lean_Meta_addCustomEliminatorEntry___spec__6(x_23, x_25, lean_box(0), x_23, x_25, x_38); +lean_dec(x_25); +lean_dec(x_23); +if (x_39 == 0) { -lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_43 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_33, x_34, x_4, x_5); -x_44 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_44, 0, x_43); -x_45 = lean_array_fset(x_17, x_12, x_44); +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +lean_dec(x_21); +x_40 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_4, x_5); +x_41 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_41, 0, x_40); +x_42 = lean_array_fset(x_18, x_12, x_41); lean_dec(x_12); -lean_ctor_set(x_1, 0, x_45); -return x_1; +if (lean_is_scalar(x_7)) { + x_43 = lean_alloc_ctor(0, 1, 0); +} else { + x_43 = x_7; +} +lean_ctor_set(x_43, 0, x_42); +return x_43; } else { -lean_object* x_46; lean_object* x_47; -lean_dec(x_34); -lean_dec(x_33); -x_46 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_46, 0, x_4); -lean_ctor_set(x_46, 1, x_5); -x_47 = lean_array_fset(x_17, x_12, x_46); +lean_object* x_44; lean_object* x_45; lean_object* x_46; +lean_dec(x_20); +lean_dec(x_19); +if (lean_is_scalar(x_21)) { + x_44 = lean_alloc_ctor(0, 2, 0); +} else { + x_44 = x_21; +} +lean_ctor_set(x_44, 0, x_4); +lean_ctor_set(x_44, 1, x_5); +x_45 = lean_array_fset(x_18, x_12, x_44); lean_dec(x_12); -lean_ctor_set(x_1, 0, x_47); -return x_1; +if (lean_is_scalar(x_7)) { + x_46 = lean_alloc_ctor(0, 1, 0); +} else { + x_46 = x_7; +} +lean_ctor_set(x_46, 0, x_45); +return x_46; +} } } } } case 1: { -uint8_t x_48; -x_48 = !lean_is_exclusive(x_15); -if (x_48 == 0) -{ -lean_object* x_49; size_t x_50; size_t x_51; lean_object* x_52; lean_object* x_53; -x_49 = lean_ctor_get(x_15, 0); -x_50 = lean_usize_shift_right(x_2, x_9); -x_51 = lean_usize_add(x_3, x_8); -x_52 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3(x_49, x_50, x_51, x_4, x_5); -lean_ctor_set(x_15, 0, x_52); -x_53 = lean_array_fset(x_17, x_12, x_15); +uint8_t x_55; +x_55 = !lean_is_exclusive(x_16); +if (x_55 == 0) +{ +lean_object* x_56; size_t x_57; size_t x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_56 = lean_ctor_get(x_16, 0); +x_57 = lean_usize_shift_right(x_2, x_9); +x_58 = lean_usize_add(x_3, x_8); +x_59 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3(x_56, x_57, x_58, x_4, x_5); +lean_ctor_set(x_16, 0, x_59); +x_60 = lean_array_fset(x_18, x_12, x_16); lean_dec(x_12); -lean_ctor_set(x_1, 0, x_53); -return x_1; +if (lean_is_scalar(x_7)) { + x_61 = lean_alloc_ctor(0, 1, 0); +} else { + x_61 = x_7; +} +lean_ctor_set(x_61, 0, x_60); +return x_61; } else { -lean_object* x_54; size_t x_55; size_t x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_54 = lean_ctor_get(x_15, 0); -lean_inc(x_54); -lean_dec(x_15); -x_55 = lean_usize_shift_right(x_2, x_9); -x_56 = lean_usize_add(x_3, x_8); -x_57 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3(x_54, x_55, x_56, x_4, x_5); -x_58 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_58, 0, x_57); -x_59 = lean_array_fset(x_17, x_12, x_58); +lean_object* x_62; size_t x_63; size_t x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_62 = lean_ctor_get(x_16, 0); +lean_inc(x_62); +lean_dec(x_16); +x_63 = lean_usize_shift_right(x_2, x_9); +x_64 = lean_usize_add(x_3, x_8); +x_65 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3(x_62, x_63, x_64, x_4, x_5); +x_66 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_66, 0, x_65); +x_67 = lean_array_fset(x_18, x_12, x_66); lean_dec(x_12); -lean_ctor_set(x_1, 0, x_59); -return x_1; +if (lean_is_scalar(x_7)) { + x_68 = lean_alloc_ctor(0, 1, 0); +} else { + x_68 = x_7; +} +lean_ctor_set(x_68, 0, x_67); +return x_68; } } default: { -lean_object* x_60; lean_object* x_61; -x_60 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_60, 0, x_4); -lean_ctor_set(x_60, 1, x_5); -x_61 = lean_array_fset(x_17, x_12, x_60); +lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_69 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_69, 0, x_4); +lean_ctor_set(x_69, 1, x_5); +x_70 = lean_array_fset(x_18, x_12, x_69); lean_dec(x_12); -lean_ctor_set(x_1, 0, x_61); -return x_1; +if (lean_is_scalar(x_7)) { + x_71 = lean_alloc_ctor(0, 1, 0); +} else { + x_71 = x_7; } +lean_ctor_set(x_71, 0, x_70); +return x_71; } } } -else -{ -lean_object* x_62; size_t x_63; size_t x_64; size_t x_65; size_t x_66; lean_object* x_67; lean_object* x_68; uint8_t x_69; -x_62 = lean_ctor_get(x_1, 0); -lean_inc(x_62); -lean_dec(x_1); -x_63 = 1; -x_64 = 5; -x_65 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3___closed__2; -x_66 = lean_usize_land(x_2, x_65); -x_67 = lean_usize_to_nat(x_66); -x_68 = lean_array_get_size(x_62); -x_69 = lean_nat_dec_lt(x_67, x_68); -lean_dec(x_68); -if (x_69 == 0) -{ -lean_object* x_70; -lean_dec(x_67); -lean_dec(x_5); -lean_dec(x_4); -x_70 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_70, 0, x_62); -return x_70; } else { -lean_object* x_71; lean_object* x_72; lean_object* x_73; -x_71 = lean_array_fget(x_62, x_67); -x_72 = lean_box(0); -x_73 = lean_array_fset(x_62, x_67, x_72); -switch (lean_obj_tag(x_71)) { -case 0: +uint8_t x_72; +x_72 = !lean_is_exclusive(x_1); +if (x_72 == 0) { -lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; uint8_t x_79; -x_74 = lean_ctor_get(x_71, 0); -lean_inc(x_74); -x_75 = lean_ctor_get(x_71, 1); -lean_inc(x_75); -if (lean_is_exclusive(x_71)) { - lean_ctor_release(x_71, 0); - lean_ctor_release(x_71, 1); - x_76 = x_71; -} else { - lean_dec_ref(x_71); - x_76 = lean_box(0); -} -x_77 = lean_array_get_size(x_4); -x_78 = lean_array_get_size(x_74); -x_79 = lean_nat_dec_eq(x_77, x_78); -lean_dec(x_78); +lean_object* x_73; lean_object* x_74; size_t x_75; uint8_t x_76; +x_73 = lean_unsigned_to_nat(0u); +x_74 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_addCustomEliminatorEntry___spec__7(x_1, x_73, x_4, x_5); +x_75 = 7; +x_76 = lean_usize_dec_le(x_75, x_3); +if (x_76 == 0) +{ +lean_object* x_77; lean_object* x_78; uint8_t x_79; +x_77 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_74); +x_78 = lean_unsigned_to_nat(4u); +x_79 = lean_nat_dec_lt(x_77, x_78); lean_dec(x_77); if (x_79 == 0) { lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; -lean_dec(x_76); -x_80 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_74, x_75, x_4, x_5); -x_81 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_81, 0, x_80); -x_82 = lean_array_fset(x_73, x_67, x_81); -lean_dec(x_67); -x_83 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_83, 0, x_82); +x_80 = lean_ctor_get(x_74, 0); +lean_inc(x_80); +x_81 = lean_ctor_get(x_74, 1); +lean_inc(x_81); +lean_dec(x_74); +x_82 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3___closed__3; +x_83 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4(x_3, x_80, x_81, lean_box(0), x_73, x_82); +lean_dec(x_81); +lean_dec(x_80); return x_83; } else { -lean_object* x_84; uint8_t x_85; -x_84 = lean_unsigned_to_nat(0u); -x_85 = l_Array_isEqvAux___at_Lean_Meta_addCustomEliminatorEntry___spec__6(x_4, x_74, lean_box(0), x_4, x_74, x_84); -if (x_85 == 0) -{ -lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; -lean_dec(x_76); -x_86 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_74, x_75, x_4, x_5); -x_87 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_87, 0, x_86); -x_88 = lean_array_fset(x_73, x_67, x_87); -lean_dec(x_67); -x_89 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_89, 0, x_88); -return x_89; +return x_74; +} } else { -lean_object* x_90; lean_object* x_91; lean_object* x_92; -lean_dec(x_75); -lean_dec(x_74); -if (lean_is_scalar(x_76)) { - x_90 = lean_alloc_ctor(0, 2, 0); -} else { - x_90 = x_76; -} -lean_ctor_set(x_90, 0, x_4); -lean_ctor_set(x_90, 1, x_5); -x_91 = lean_array_fset(x_73, x_67, x_90); -lean_dec(x_67); -x_92 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_92, 0, x_91); -return x_92; +return x_74; } } +else +{ +lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; size_t x_89; uint8_t x_90; +x_84 = lean_ctor_get(x_1, 0); +x_85 = lean_ctor_get(x_1, 1); +lean_inc(x_85); +lean_inc(x_84); +lean_dec(x_1); +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_87 = lean_unsigned_to_nat(0u); +x_88 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_addCustomEliminatorEntry___spec__7(x_86, x_87, x_4, x_5); +x_89 = 7; +x_90 = lean_usize_dec_le(x_89, x_3); +if (x_90 == 0) +{ +lean_object* x_91; lean_object* x_92; uint8_t x_93; +x_91 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_88); +x_92 = lean_unsigned_to_nat(4u); +x_93 = lean_nat_dec_lt(x_91, x_92); +lean_dec(x_91); +if (x_93 == 0) +{ +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; +x_94 = lean_ctor_get(x_88, 0); +lean_inc(x_94); +x_95 = lean_ctor_get(x_88, 1); +lean_inc(x_95); +lean_dec(x_88); +x_96 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3___closed__3; +x_97 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4(x_3, x_94, x_95, lean_box(0), x_87, x_96); +lean_dec(x_95); +lean_dec(x_94); +return x_97; } -case 1: +else { -lean_object* x_93; lean_object* x_94; size_t x_95; size_t x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; -x_93 = lean_ctor_get(x_71, 0); -lean_inc(x_93); -if (lean_is_exclusive(x_71)) { - lean_ctor_release(x_71, 0); - x_94 = x_71; -} else { - lean_dec_ref(x_71); - x_94 = lean_box(0); -} -x_95 = lean_usize_shift_right(x_2, x_64); -x_96 = lean_usize_add(x_3, x_63); -x_97 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3(x_93, x_95, x_96, x_4, x_5); -if (lean_is_scalar(x_94)) { - x_98 = lean_alloc_ctor(1, 1, 0); -} else { - x_98 = x_94; +return x_88; } -lean_ctor_set(x_98, 0, x_97); -x_99 = lean_array_fset(x_73, x_67, x_98); -lean_dec(x_67); -x_100 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_100, 0, x_99); -return x_100; } -default: +else { -lean_object* x_101; lean_object* x_102; lean_object* x_103; -x_101 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_101, 0, x_4); -lean_ctor_set(x_101, 1, x_5); -x_102 = lean_array_fset(x_73, x_67, x_101); -lean_dec(x_67); -x_103 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_103, 0, x_102); -return x_103; +return x_88; } } } } } -else +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_addCustomEliminatorEntry___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -uint8_t x_104; -x_104 = !lean_is_exclusive(x_1); -if (x_104 == 0) +uint8_t x_4; +x_4 = !lean_is_exclusive(x_1); +if (x_4 == 0) { -lean_object* x_105; lean_object* x_106; size_t x_107; uint8_t x_108; -x_105 = lean_unsigned_to_nat(0u); -x_106 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_addCustomEliminatorEntry___spec__7(x_1, x_105, x_4, x_5); -x_107 = 7; -x_108 = lean_usize_dec_le(x_107, x_3); -if (x_108 == 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; uint8_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_11); +x_14 = lean_unsigned_to_nat(0u); +x_15 = lean_nat_dec_lt(x_14, x_13); +x_16 = lean_unbox(x_10); +lean_dec(x_10); +if (x_16 == 0) { -lean_object* x_109; lean_object* x_110; uint8_t x_111; -x_109 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_106); -x_110 = lean_unsigned_to_nat(4u); -x_111 = lean_nat_dec_lt(x_109, x_110); -lean_dec(x_109); -if (x_111 == 0) +if (x_15 == 0) { -lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; -x_112 = lean_ctor_get(x_106, 0); -lean_inc(x_112); -x_113 = lean_ctor_get(x_106, 1); -lean_inc(x_113); -lean_dec(x_106); -x_114 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3___closed__3; -x_115 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4(x_3, x_112, x_113, lean_box(0), x_105, x_114); -lean_dec(x_113); -lean_dec(x_112); -return x_115; +size_t x_17; lean_object* x_18; +lean_dec(x_13); +lean_dec(x_11); +x_17 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__2; +x_18 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3(x_5, x_17, x_7, x_2, x_3); +lean_ctor_set(x_1, 1, x_9); +lean_ctor_set(x_1, 0, x_18); +return x_1; } else { -return x_106; -} +uint64_t x_19; uint8_t x_20; +x_19 = 13; +x_20 = lean_nat_dec_le(x_13, x_13); +if (x_20 == 0) +{ +size_t x_21; lean_object* x_22; +lean_dec(x_13); +lean_dec(x_11); +x_21 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__2; +x_22 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3(x_5, x_21, x_7, x_2, x_3); +lean_ctor_set(x_1, 1, x_9); +lean_ctor_set(x_1, 0, x_22); +return x_1; } else { -return x_106; +size_t x_23; size_t x_24; uint64_t x_25; uint64_t x_26; size_t x_27; lean_object* x_28; +x_23 = 0; +x_24 = lean_usize_of_nat(x_13); +lean_dec(x_13); +x_25 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_addCustomEliminatorEntry___spec__5(x_11, x_23, x_24, x_12); +lean_dec(x_11); +x_26 = lean_uint64_mix_hash(x_19, x_25); +x_27 = lean_uint64_to_usize(x_26); +x_28 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3(x_5, x_27, x_7, x_2, x_3); +lean_ctor_set(x_1, 1, x_9); +lean_ctor_set(x_1, 0, x_28); +return x_1; +} } } else { -lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; size_t x_121; uint8_t x_122; -x_116 = lean_ctor_get(x_1, 0); -x_117 = lean_ctor_get(x_1, 1); -lean_inc(x_117); -lean_inc(x_116); -lean_dec(x_1); -x_118 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_118, 0, x_116); -lean_ctor_set(x_118, 1, x_117); -x_119 = lean_unsigned_to_nat(0u); -x_120 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_addCustomEliminatorEntry___spec__7(x_118, x_119, x_4, x_5); -x_121 = 7; -x_122 = lean_usize_dec_le(x_121, x_3); -if (x_122 == 0) -{ -lean_object* x_123; lean_object* x_124; uint8_t x_125; -x_123 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_120); -x_124 = lean_unsigned_to_nat(4u); -x_125 = lean_nat_dec_lt(x_123, x_124); -lean_dec(x_123); -if (x_125 == 0) -{ -lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; -x_126 = lean_ctor_get(x_120, 0); -lean_inc(x_126); -x_127 = lean_ctor_get(x_120, 1); -lean_inc(x_127); -lean_dec(x_120); -x_128 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3___closed__3; -x_129 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4(x_3, x_126, x_127, lean_box(0), x_119, x_128); -lean_dec(x_127); -lean_dec(x_126); -return x_129; +if (x_15 == 0) +{ +size_t x_29; lean_object* x_30; +lean_dec(x_13); +lean_dec(x_11); +x_29 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__4; +x_30 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3(x_5, x_29, x_7, x_2, x_3); +lean_ctor_set(x_1, 1, x_9); +lean_ctor_set(x_1, 0, x_30); +return x_1; } else { -return x_120; -} +uint64_t x_31; uint8_t x_32; +x_31 = 11; +x_32 = lean_nat_dec_le(x_13, x_13); +if (x_32 == 0) +{ +size_t x_33; lean_object* x_34; +lean_dec(x_13); +lean_dec(x_11); +x_33 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__4; +x_34 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3(x_5, x_33, x_7, x_2, x_3); +lean_ctor_set(x_1, 1, x_9); +lean_ctor_set(x_1, 0, x_34); +return x_1; } else { -return x_120; -} +size_t x_35; size_t x_36; uint64_t x_37; uint64_t x_38; size_t x_39; lean_object* x_40; +x_35 = 0; +x_36 = lean_usize_of_nat(x_13); +lean_dec(x_13); +x_37 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_addCustomEliminatorEntry___spec__5(x_11, x_35, x_36, x_12); +lean_dec(x_11); +x_38 = lean_uint64_mix_hash(x_31, x_37); +x_39 = lean_uint64_to_usize(x_38); +x_40 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3(x_5, x_39, x_7, x_2, x_3); +lean_ctor_set(x_1, 1, x_9); +lean_ctor_set(x_1, 0, x_40); +return x_1; } } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_addCustomEliminatorEntry___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -uint8_t x_4; -x_4 = !lean_is_exclusive(x_1); -if (x_4 == 0) +lean_object* x_41; lean_object* x_42; size_t x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint64_t x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; uint8_t x_52; +x_41 = lean_ctor_get(x_1, 0); +x_42 = lean_ctor_get(x_1, 1); +lean_inc(x_42); +lean_inc(x_41); +lean_dec(x_1); +x_43 = 1; +x_44 = lean_unsigned_to_nat(1u); +x_45 = lean_nat_add(x_42, x_44); +lean_dec(x_42); +x_46 = lean_ctor_get(x_2, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_2, 1); +lean_inc(x_47); +x_48 = 7; +x_49 = lean_array_get_size(x_47); +x_50 = lean_unsigned_to_nat(0u); +x_51 = lean_nat_dec_lt(x_50, x_49); +x_52 = lean_unbox(x_46); +lean_dec(x_46); +if (x_52 == 0) { -lean_object* x_5; lean_object* x_6; uint64_t x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; -x_5 = lean_ctor_get(x_1, 0); -x_6 = lean_ctor_get(x_1, 1); -x_7 = 7; -x_8 = lean_array_get_size(x_2); -x_9 = lean_unsigned_to_nat(0u); -x_10 = lean_nat_dec_lt(x_9, x_8); -x_11 = 1; -x_12 = lean_unsigned_to_nat(1u); -x_13 = lean_nat_add(x_6, x_12); -lean_dec(x_6); -if (x_10 == 0) +if (x_51 == 0) { -size_t x_14; lean_object* x_15; -lean_dec(x_8); -x_14 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__1; -x_15 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3(x_5, x_14, x_11, x_2, x_3); -lean_ctor_set(x_1, 1, x_13); -lean_ctor_set(x_1, 0, x_15); -return x_1; +size_t x_53; lean_object* x_54; lean_object* x_55; +lean_dec(x_49); +lean_dec(x_47); +x_53 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__2; +x_54 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3(x_41, x_53, x_43, x_2, x_3); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_45); +return x_55; } else { -uint8_t x_16; -x_16 = lean_nat_dec_le(x_8, x_8); -if (x_16 == 0) +uint64_t x_56; uint8_t x_57; +x_56 = 13; +x_57 = lean_nat_dec_le(x_49, x_49); +if (x_57 == 0) { -size_t x_17; lean_object* x_18; -lean_dec(x_8); -x_17 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__1; -x_18 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3(x_5, x_17, x_11, x_2, x_3); -lean_ctor_set(x_1, 1, x_13); -lean_ctor_set(x_1, 0, x_18); -return x_1; +size_t x_58; lean_object* x_59; lean_object* x_60; +lean_dec(x_49); +lean_dec(x_47); +x_58 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__2; +x_59 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3(x_41, x_58, x_43, x_2, x_3); +x_60 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_60, 0, x_59); +lean_ctor_set(x_60, 1, x_45); +return x_60; } else { -size_t x_19; size_t x_20; uint64_t x_21; size_t x_22; lean_object* x_23; -x_19 = 0; -x_20 = lean_usize_of_nat(x_8); -lean_dec(x_8); -x_21 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_addCustomEliminatorEntry___spec__5(x_2, x_19, x_20, x_7); -x_22 = lean_uint64_to_usize(x_21); -x_23 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3(x_5, x_22, x_11, x_2, x_3); -lean_ctor_set(x_1, 1, x_13); -lean_ctor_set(x_1, 0, x_23); -return x_1; +size_t x_61; size_t x_62; uint64_t x_63; uint64_t x_64; size_t x_65; lean_object* x_66; lean_object* x_67; +x_61 = 0; +x_62 = lean_usize_of_nat(x_49); +lean_dec(x_49); +x_63 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_addCustomEliminatorEntry___spec__5(x_47, x_61, x_62, x_48); +lean_dec(x_47); +x_64 = lean_uint64_mix_hash(x_56, x_63); +x_65 = lean_uint64_to_usize(x_64); +x_66 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3(x_41, x_65, x_43, x_2, x_3); +x_67 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_67, 0, x_66); +lean_ctor_set(x_67, 1, x_45); +return x_67; } } } else { -lean_object* x_24; lean_object* x_25; uint64_t x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; size_t x_30; lean_object* x_31; lean_object* x_32; -x_24 = lean_ctor_get(x_1, 0); -x_25 = lean_ctor_get(x_1, 1); -lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_1); -x_26 = 7; -x_27 = lean_array_get_size(x_2); -x_28 = lean_unsigned_to_nat(0u); -x_29 = lean_nat_dec_lt(x_28, x_27); -x_30 = 1; -x_31 = lean_unsigned_to_nat(1u); -x_32 = lean_nat_add(x_25, x_31); -lean_dec(x_25); -if (x_29 == 0) +if (x_51 == 0) { -size_t x_33; lean_object* x_34; lean_object* x_35; -lean_dec(x_27); -x_33 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__1; -x_34 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3(x_24, x_33, x_30, x_2, x_3); -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; +size_t x_68; lean_object* x_69; lean_object* x_70; +lean_dec(x_49); +lean_dec(x_47); +x_68 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__4; +x_69 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3(x_41, x_68, x_43, x_2, x_3); +x_70 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_70, 0, x_69); +lean_ctor_set(x_70, 1, x_45); +return x_70; } else { -uint8_t x_36; -x_36 = lean_nat_dec_le(x_27, x_27); -if (x_36 == 0) +uint64_t x_71; uint8_t x_72; +x_71 = 11; +x_72 = lean_nat_dec_le(x_49, x_49); +if (x_72 == 0) { -size_t x_37; lean_object* x_38; lean_object* x_39; -lean_dec(x_27); -x_37 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__1; -x_38 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3(x_24, x_37, x_30, x_2, x_3); -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_32); -return x_39; +size_t x_73; lean_object* x_74; lean_object* x_75; +lean_dec(x_49); +lean_dec(x_47); +x_73 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__4; +x_74 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3(x_41, x_73, x_43, x_2, x_3); +x_75 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_75, 0, x_74); +lean_ctor_set(x_75, 1, x_45); +return x_75; } else { -size_t x_40; size_t x_41; uint64_t x_42; size_t x_43; lean_object* x_44; lean_object* x_45; -x_40 = 0; -x_41 = lean_usize_of_nat(x_27); -lean_dec(x_27); -x_42 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_addCustomEliminatorEntry___spec__5(x_2, x_40, x_41, x_26); -x_43 = lean_uint64_to_usize(x_42); -x_44 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3(x_24, x_43, x_30, 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_32); -return x_45; +size_t x_76; size_t x_77; uint64_t x_78; uint64_t x_79; size_t x_80; lean_object* x_81; lean_object* x_82; +x_76 = 0; +x_77 = lean_usize_of_nat(x_49); +lean_dec(x_49); +x_78 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_addCustomEliminatorEntry___spec__5(x_47, x_76, x_77, x_48); +lean_dec(x_47); +x_79 = lean_uint64_mix_hash(x_71, x_78); +x_80 = lean_uint64_to_usize(x_79); +x_81 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3(x_41, x_80, x_43, x_2, x_3); +x_82 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_82, 0, x_81); +lean_ctor_set(x_82, 1, x_45); +return x_82; +} } } } @@ -6611,34 +7006,89 @@ return x_3; } else { -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; uint8_t x_21; x_4 = lean_ctor_get(x_2, 0); x_5 = lean_ctor_get(x_2, 2); -x_6 = lean_array_get_size(x_4); -x_7 = lean_array_get_size(x_1); -x_8 = lean_nat_dec_eq(x_6, x_7); -lean_dec(x_7); -lean_dec(x_6); -if (x_8 == 0) +x_6 = lean_ctor_get(x_4, 0); +x_7 = lean_ctor_get(x_4, 1); +x_8 = lean_ctor_get(x_1, 0); +x_9 = lean_ctor_get(x_1, 1); +x_21 = lean_unbox(x_6); +if (x_21 == 0) +{ +uint8_t x_22; +x_22 = lean_unbox(x_8); +if (x_22 == 0) +{ +uint8_t x_23; +x_23 = 1; +x_10 = x_23; +goto block_20; +} +else +{ +uint8_t x_24; +x_24 = 0; +x_10 = x_24; +goto block_20; +} +} +else +{ +uint8_t x_25; +x_25 = lean_unbox(x_8); +if (x_25 == 0) +{ +uint8_t x_26; +x_26 = 0; +x_10 = x_26; +goto block_20; +} +else +{ +uint8_t x_27; +x_27 = 1; +x_10 = x_27; +goto block_20; +} +} +block_20: +{ +if (x_10 == 0) { x_2 = x_5; goto _start; } else { -lean_object* x_10; uint8_t x_11; -x_10 = lean_unsigned_to_nat(0u); -x_11 = l_Array_isEqvAux___at_Lean_Meta_addCustomEliminatorEntry___spec__11(x_1, x_4, lean_box(0), x_4, x_1, x_10); -if (x_11 == 0) +lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_12 = lean_array_get_size(x_7); +x_13 = lean_array_get_size(x_9); +x_14 = lean_nat_dec_eq(x_12, x_13); +lean_dec(x_13); +lean_dec(x_12); +if (x_14 == 0) { x_2 = x_5; goto _start; } else { -uint8_t x_13; -x_13 = 1; -return x_13; +lean_object* x_16; uint8_t x_17; +x_16 = lean_unsigned_to_nat(0u); +x_17 = l_Array_isEqvAux___at_Lean_Meta_addCustomEliminatorEntry___spec__11(x_7, x_9, lean_box(0), x_7, x_9, x_16); +if (x_17 == 0) +{ +x_2 = x_5; +goto _start; +} +else +{ +uint8_t x_19; +x_19 = 1; +return x_19; +} +} } } } @@ -6714,131 +7164,128 @@ return x_1; } else { -uint8_t x_3; -x_3 = !lean_is_exclusive(x_2); -if (x_3 == 0) -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; uint64_t x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; -x_4 = lean_ctor_get(x_2, 0); +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint64_t x_8; lean_object* x_15; lean_object* x_16; uint64_t x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; uint8_t x_21; +x_3 = lean_ctor_get(x_2, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_2, 1); +lean_inc(x_4); x_5 = lean_ctor_get(x_2, 2); -x_6 = lean_array_get_size(x_1); -x_7 = 7; -x_8 = lean_array_get_size(x_4); -x_9 = lean_unsigned_to_nat(0u); -x_10 = lean_nat_dec_lt(x_9, x_8); -if (x_10 == 0) +lean_inc(x_5); +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_6 = x_2; +} else { + lean_dec_ref(x_2); + x_6 = lean_box(0); +} +x_7 = lean_array_get_size(x_1); +x_15 = lean_ctor_get(x_3, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_3, 1); +lean_inc(x_16); +x_17 = 7; +x_18 = lean_array_get_size(x_16); +x_19 = lean_unsigned_to_nat(0u); +x_20 = lean_nat_dec_lt(x_19, x_18); +x_21 = lean_unbox(x_15); +lean_dec(x_15); +if (x_21 == 0) { -size_t x_11; lean_object* x_12; lean_object* x_13; -lean_dec(x_8); -x_11 = lean_hashmap_mk_idx(x_6, x_7); -x_12 = lean_array_uget(x_1, x_11); -lean_ctor_set(x_2, 2, x_12); -x_13 = lean_array_uset(x_1, x_11, x_2); -x_1 = x_13; -x_2 = x_5; -goto _start; +if (x_20 == 0) +{ +uint64_t x_22; +lean_dec(x_18); +lean_dec(x_16); +x_22 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__1; +x_8 = x_22; +goto block_14; } else { -uint8_t x_15; -x_15 = lean_nat_dec_le(x_8, x_8); -if (x_15 == 0) +uint64_t x_23; uint8_t x_24; +x_23 = 13; +x_24 = lean_nat_dec_le(x_18, x_18); +if (x_24 == 0) { -size_t x_16; lean_object* x_17; lean_object* x_18; -lean_dec(x_8); -x_16 = lean_hashmap_mk_idx(x_6, x_7); -x_17 = lean_array_uget(x_1, x_16); -lean_ctor_set(x_2, 2, x_17); -x_18 = lean_array_uset(x_1, x_16, x_2); -x_1 = x_18; -x_2 = x_5; -goto _start; +uint64_t x_25; +lean_dec(x_18); +lean_dec(x_16); +x_25 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__1; +x_8 = x_25; +goto block_14; } else { -size_t x_20; size_t x_21; uint64_t x_22; size_t x_23; lean_object* x_24; lean_object* x_25; -x_20 = 0; -x_21 = lean_usize_of_nat(x_8); -lean_dec(x_8); -x_22 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_addCustomEliminatorEntry___spec__5(x_4, x_20, x_21, x_7); -x_23 = lean_hashmap_mk_idx(x_6, x_22); -x_24 = lean_array_uget(x_1, x_23); -lean_ctor_set(x_2, 2, x_24); -x_25 = lean_array_uset(x_1, x_23, x_2); -x_1 = x_25; -x_2 = x_5; -goto _start; +size_t x_26; size_t x_27; uint64_t x_28; uint64_t x_29; +x_26 = 0; +x_27 = lean_usize_of_nat(x_18); +lean_dec(x_18); +x_28 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_addCustomEliminatorEntry___spec__5(x_16, x_26, x_27, x_17); +lean_dec(x_16); +x_29 = lean_uint64_mix_hash(x_23, x_28); +x_8 = x_29; +goto block_14; } } } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint64_t x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; -x_27 = lean_ctor_get(x_2, 0); -x_28 = lean_ctor_get(x_2, 1); -x_29 = lean_ctor_get(x_2, 2); -lean_inc(x_29); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_2); -x_30 = lean_array_get_size(x_1); -x_31 = 7; -x_32 = lean_array_get_size(x_27); -x_33 = lean_unsigned_to_nat(0u); -x_34 = lean_nat_dec_lt(x_33, x_32); -if (x_34 == 0) +if (x_20 == 0) { -size_t x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -lean_dec(x_32); -x_35 = lean_hashmap_mk_idx(x_30, x_31); -x_36 = lean_array_uget(x_1, x_35); -x_37 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_37, 0, x_27); -lean_ctor_set(x_37, 1, x_28); -lean_ctor_set(x_37, 2, x_36); -x_38 = lean_array_uset(x_1, x_35, x_37); -x_1 = x_38; -x_2 = x_29; -goto _start; +uint64_t x_30; +lean_dec(x_18); +lean_dec(x_16); +x_30 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__3; +x_8 = x_30; +goto block_14; } else { -uint8_t x_40; -x_40 = lean_nat_dec_le(x_32, x_32); -if (x_40 == 0) +uint64_t x_31; uint8_t x_32; +x_31 = 11; +x_32 = lean_nat_dec_le(x_18, x_18); +if (x_32 == 0) { -size_t x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; -lean_dec(x_32); -x_41 = lean_hashmap_mk_idx(x_30, x_31); -x_42 = lean_array_uget(x_1, x_41); -x_43 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_43, 0, x_27); -lean_ctor_set(x_43, 1, x_28); -lean_ctor_set(x_43, 2, x_42); -x_44 = lean_array_uset(x_1, x_41, x_43); -x_1 = x_44; -x_2 = x_29; -goto _start; +uint64_t x_33; +lean_dec(x_18); +lean_dec(x_16); +x_33 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__3; +x_8 = x_33; +goto block_14; } else { -size_t x_46; size_t x_47; uint64_t x_48; size_t x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_46 = 0; -x_47 = lean_usize_of_nat(x_32); -lean_dec(x_32); -x_48 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_addCustomEliminatorEntry___spec__5(x_27, x_46, x_47, x_31); -x_49 = lean_hashmap_mk_idx(x_30, x_48); -x_50 = lean_array_uget(x_1, x_49); -x_51 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_51, 0, x_27); -lean_ctor_set(x_51, 1, x_28); -lean_ctor_set(x_51, 2, x_50); -x_52 = lean_array_uset(x_1, x_49, x_51); -x_1 = x_52; -x_2 = x_29; -goto _start; +size_t x_34; size_t x_35; uint64_t x_36; uint64_t x_37; +x_34 = 0; +x_35 = lean_usize_of_nat(x_18); +lean_dec(x_18); +x_36 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_addCustomEliminatorEntry___spec__5(x_16, x_34, x_35, x_17); +lean_dec(x_16); +x_37 = lean_uint64_mix_hash(x_31, x_36); +x_8 = x_37; +goto block_14; +} } } +block_14: +{ +size_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_hashmap_mk_idx(x_7, x_8); +x_10 = lean_array_uget(x_1, x_9); +if (lean_is_scalar(x_6)) { + x_11 = lean_alloc_ctor(1, 3, 0); +} else { + x_11 = x_6; +} +lean_ctor_set(x_11, 0, x_3); +lean_ctor_set(x_11, 1, x_4); +lean_ctor_set(x_11, 2, x_10); +x_12 = lean_array_uset(x_1, x_9, x_11); +x_1 = x_12; +x_2 = x_5; +goto _start; } } } @@ -6946,98 +7393,150 @@ return x_4; } else { -uint8_t x_5; -x_5 = !lean_is_exclusive(x_3); -if (x_5 == 0) -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_6 = lean_ctor_get(x_3, 0); -x_7 = lean_ctor_get(x_3, 1); -x_8 = lean_ctor_get(x_3, 2); -x_9 = lean_array_get_size(x_6); -x_10 = lean_array_get_size(x_1); -x_11 = lean_nat_dec_eq(x_9, x_10); -lean_dec(x_10); +lean_object* x_5; lean_object* x_6; 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_27; +x_5 = lean_ctor_get(x_3, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_3, 1); +lean_inc(x_6); +x_7 = lean_ctor_get(x_3, 2); +lean_inc(x_7); +if (lean_is_exclusive(x_3)) { + lean_ctor_release(x_3, 0); + lean_ctor_release(x_3, 1); + lean_ctor_release(x_3, 2); + x_8 = x_3; +} else { + lean_dec_ref(x_3); + x_8 = lean_box(0); +} +x_9 = lean_ctor_get(x_5, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_5, 1); +lean_inc(x_10); +x_11 = lean_ctor_get(x_1, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_1, 1); +lean_inc(x_12); +x_27 = lean_unbox(x_9); lean_dec(x_9); -if (x_11 == 0) +if (x_27 == 0) { -lean_object* x_12; -x_12 = l_Lean_AssocList_replace___at_Lean_Meta_addCustomEliminatorEntry___spec__16(x_1, x_2, x_8); -lean_ctor_set(x_3, 2, x_12); -return x_3; +uint8_t x_28; +x_28 = lean_unbox(x_11); +lean_dec(x_11); +if (x_28 == 0) +{ +uint8_t x_29; +x_29 = 1; +x_13 = x_29; +goto block_26; } else { -lean_object* x_13; uint8_t x_14; -x_13 = lean_unsigned_to_nat(0u); -x_14 = l_Array_isEqvAux___at_Lean_Meta_addCustomEliminatorEntry___spec__17(x_1, x_6, lean_box(0), x_6, x_1, x_13); -if (x_14 == 0) +uint8_t x_30; +x_30 = 0; +x_13 = x_30; +goto block_26; +} +} +else { -lean_object* x_15; -x_15 = l_Lean_AssocList_replace___at_Lean_Meta_addCustomEliminatorEntry___spec__16(x_1, x_2, x_8); -lean_ctor_set(x_3, 2, x_15); -return x_3; +uint8_t x_31; +x_31 = lean_unbox(x_11); +lean_dec(x_11); +if (x_31 == 0) +{ +uint8_t x_32; +x_32 = 0; +x_13 = x_32; +goto block_26; } else { -lean_dec(x_7); -lean_dec(x_6); -lean_ctor_set(x_3, 1, x_2); -lean_ctor_set(x_3, 0, x_1); -return x_3; +uint8_t x_33; +x_33 = 1; +x_13 = x_33; +goto block_26; } } +block_26: +{ +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; +lean_dec(x_12); +lean_dec(x_10); +x_14 = l_Lean_AssocList_replace___at_Lean_Meta_addCustomEliminatorEntry___spec__16(x_1, x_2, x_7); +if (lean_is_scalar(x_8)) { + x_15 = lean_alloc_ctor(1, 3, 0); +} else { + x_15 = x_8; +} +lean_ctor_set(x_15, 0, x_5); +lean_ctor_set(x_15, 1, x_6); +lean_ctor_set(x_15, 2, x_14); +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_3, 0); -x_17 = lean_ctor_get(x_3, 1); -x_18 = lean_ctor_get(x_3, 2); -lean_inc(x_18); -lean_inc(x_17); -lean_inc(x_16); -lean_dec(x_3); -x_19 = lean_array_get_size(x_16); -x_20 = lean_array_get_size(x_1); -x_21 = lean_nat_dec_eq(x_19, x_20); -lean_dec(x_20); -lean_dec(x_19); -if (x_21 == 0) +lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_16 = lean_array_get_size(x_10); +x_17 = lean_array_get_size(x_12); +x_18 = lean_nat_dec_eq(x_16, x_17); +lean_dec(x_17); +lean_dec(x_16); +if (x_18 == 0) { -lean_object* x_22; lean_object* x_23; -x_22 = l_Lean_AssocList_replace___at_Lean_Meta_addCustomEliminatorEntry___spec__16(x_1, x_2, x_18); -x_23 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_23, 0, x_16); -lean_ctor_set(x_23, 1, x_17); -lean_ctor_set(x_23, 2, x_22); -return x_23; +lean_object* x_19; lean_object* x_20; +lean_dec(x_12); +lean_dec(x_10); +x_19 = l_Lean_AssocList_replace___at_Lean_Meta_addCustomEliminatorEntry___spec__16(x_1, x_2, x_7); +if (lean_is_scalar(x_8)) { + x_20 = lean_alloc_ctor(1, 3, 0); +} else { + x_20 = x_8; +} +lean_ctor_set(x_20, 0, x_5); +lean_ctor_set(x_20, 1, x_6); +lean_ctor_set(x_20, 2, x_19); +return x_20; } else { -lean_object* x_24; uint8_t x_25; -x_24 = lean_unsigned_to_nat(0u); -x_25 = l_Array_isEqvAux___at_Lean_Meta_addCustomEliminatorEntry___spec__17(x_1, x_16, lean_box(0), x_16, x_1, x_24); -if (x_25 == 0) +lean_object* x_21; uint8_t x_22; +x_21 = lean_unsigned_to_nat(0u); +x_22 = l_Array_isEqvAux___at_Lean_Meta_addCustomEliminatorEntry___spec__17(x_10, x_12, lean_box(0), x_10, x_12, x_21); +lean_dec(x_12); +lean_dec(x_10); +if (x_22 == 0) { -lean_object* x_26; lean_object* x_27; -x_26 = l_Lean_AssocList_replace___at_Lean_Meta_addCustomEliminatorEntry___spec__16(x_1, x_2, x_18); -x_27 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_27, 0, x_16); -lean_ctor_set(x_27, 1, x_17); -lean_ctor_set(x_27, 2, x_26); -return x_27; +lean_object* x_23; lean_object* x_24; +x_23 = l_Lean_AssocList_replace___at_Lean_Meta_addCustomEliminatorEntry___spec__16(x_1, x_2, x_7); +if (lean_is_scalar(x_8)) { + x_24 = lean_alloc_ctor(1, 3, 0); +} else { + x_24 = x_8; +} +lean_ctor_set(x_24, 0, x_5); +lean_ctor_set(x_24, 1, x_6); +lean_ctor_set(x_24, 2, x_23); +return x_24; } else { -lean_object* x_28; -lean_dec(x_17); -lean_dec(x_16); -x_28 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_28, 0, x_1); -lean_ctor_set(x_28, 1, x_2); -lean_ctor_set(x_28, 2, x_18); -return x_28; +lean_object* x_25; +lean_dec(x_6); +lean_dec(x_5); +if (lean_is_scalar(x_8)) { + x_25 = lean_alloc_ctor(1, 3, 0); +} else { + x_25 = x_8; +} +lean_ctor_set(x_25, 0, x_1); +lean_ctor_set(x_25, 1, x_2); +lean_ctor_set(x_25, 2, x_7); +return x_25; +} } } } @@ -7047,7 +7546,7 @@ return x_28; LEAN_EXPORT lean_object* l_Lean_HashMap_insert___at_Lean_Meta_addCustomEliminatorEntry___spec__9(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; uint64_t x_8; uint64_t x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint64_t x_8; lean_object* x_24; lean_object* x_25; uint64_t x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; uint8_t x_30; x_4 = lean_ctor_get(x_1, 0); lean_inc(x_4); x_5 = lean_ctor_get(x_1, 1); @@ -7061,37 +7560,94 @@ if (lean_is_exclusive(x_1)) { x_6 = lean_box(0); } x_7 = lean_array_get_size(x_5); -x_24 = 7; -x_25 = lean_array_get_size(x_2); -x_26 = lean_unsigned_to_nat(0u); -x_27 = lean_nat_dec_lt(x_26, x_25); -if (x_27 == 0) +x_24 = lean_ctor_get(x_2, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_2, 1); +lean_inc(x_25); +x_26 = 7; +x_27 = lean_array_get_size(x_25); +x_28 = lean_unsigned_to_nat(0u); +x_29 = lean_nat_dec_lt(x_28, x_27); +x_30 = lean_unbox(x_24); +lean_dec(x_24); +if (x_30 == 0) { +if (x_29 == 0) +{ +uint64_t x_31; +lean_dec(x_27); lean_dec(x_25); -x_8 = x_24; +x_31 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__1; +x_8 = x_31; goto block_23; } else { -uint8_t x_28; -x_28 = lean_nat_dec_le(x_25, x_25); -if (x_28 == 0) +uint64_t x_32; uint8_t x_33; +x_32 = 13; +x_33 = lean_nat_dec_le(x_27, x_27); +if (x_33 == 0) { +uint64_t x_34; +lean_dec(x_27); lean_dec(x_25); -x_8 = x_24; +x_34 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__1; +x_8 = x_34; goto block_23; } else { -size_t x_29; size_t x_30; uint64_t x_31; -x_29 = 0; -x_30 = lean_usize_of_nat(x_25); +size_t x_35; size_t x_36; uint64_t x_37; uint64_t x_38; +x_35 = 0; +x_36 = lean_usize_of_nat(x_27); +lean_dec(x_27); +x_37 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_addCustomEliminatorEntry___spec__5(x_25, x_35, x_36, x_26); lean_dec(x_25); -x_31 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_addCustomEliminatorEntry___spec__5(x_2, x_29, x_30, x_24); -x_8 = x_31; +x_38 = lean_uint64_mix_hash(x_32, x_37); +x_8 = x_38; +goto block_23; +} +} +} +else +{ +if (x_29 == 0) +{ +uint64_t x_39; +lean_dec(x_27); +lean_dec(x_25); +x_39 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__3; +x_8 = x_39; +goto block_23; +} +else +{ +uint64_t x_40; uint8_t x_41; +x_40 = 11; +x_41 = lean_nat_dec_le(x_27, x_27); +if (x_41 == 0) +{ +uint64_t x_42; +lean_dec(x_27); +lean_dec(x_25); +x_42 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__3; +x_8 = x_42; +goto block_23; +} +else +{ +size_t x_43; size_t x_44; uint64_t x_45; uint64_t x_46; +x_43 = 0; +x_44 = lean_usize_of_nat(x_27); +lean_dec(x_27); +x_45 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_addCustomEliminatorEntry___spec__5(x_25, x_43, x_44, x_26); +lean_dec(x_25); +x_46 = lean_uint64_mix_hash(x_40, x_45); +x_8 = x_46; goto block_23; } } +} block_23: { size_t x_9; lean_object* x_10; uint8_t x_11; @@ -7224,14 +7780,19 @@ return x_22; LEAN_EXPORT lean_object* l_Lean_Meta_addCustomEliminatorEntry(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); -lean_inc(x_3); -x_4 = lean_ctor_get(x_2, 1); +uint8_t x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_3 = lean_ctor_get_uint8(x_2, sizeof(void*)*2); +x_4 = lean_ctor_get(x_2, 0); lean_inc(x_4); +x_5 = lean_ctor_get(x_2, 1); +lean_inc(x_5); lean_dec(x_2); -x_5 = l_Lean_SMap_insert___at_Lean_Meta_addCustomEliminatorEntry___spec__1(x_1, x_3, x_4); -return x_5; +x_6 = lean_box(x_3); +x_7 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_4); +x_8 = l_Lean_SMap_insert___at_Lean_Meta_addCustomEliminatorEntry___spec__1(x_1, x_7, x_5); +return x_8; } } LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_addCustomEliminatorEntry___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { @@ -7275,18 +7836,17 @@ x_8 = lean_box(x_7); return x_8; } } -LEAN_EXPORT lean_object* l_Array_isEqvAux___at_Lean_Meta_addCustomEliminatorEntry___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_EXPORT lean_object* l_Array_isEqvAux___at_Lean_Meta_addCustomEliminatorEntry___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) { _start: { -uint8_t x_9; lean_object* x_10; -x_9 = l_Array_isEqvAux___at_Lean_Meta_addCustomEliminatorEntry___spec__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_3); +uint8_t x_7; lean_object* x_8; +x_7 = l_Array_isEqvAux___at_Lean_Meta_addCustomEliminatorEntry___spec__8(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_10 = lean_box(x_9); -return x_10; +x_8 = lean_box(x_7); +return x_8; } } LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { @@ -7338,7 +7898,7 @@ x_8 = lean_box(x_7); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_SMap_switch___at_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____spec__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_SMap_switch___at_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____spec__1(lean_object* x_1) { _start: { uint8_t x_2; @@ -7376,7 +7936,7 @@ return x_8; } } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__1() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____closed__1() { _start: { lean_object* x_1; @@ -7384,7 +7944,7 @@ x_1 = lean_mk_string_from_bytes("Lean", 4); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__2() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____closed__2() { _start: { lean_object* x_1; @@ -7392,7 +7952,7 @@ x_1 = lean_mk_string_from_bytes("Meta", 4); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__3() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____closed__3() { _start: { lean_object* x_1; @@ -7400,18 +7960,18 @@ x_1 = lean_mk_string_from_bytes("customEliminatorExt", 19); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__4() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____closed__4() { _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_Tactic_ElimInfo___hyg_2298____closed__1; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__2; -x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__3; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____closed__1; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____closed__2; +x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____closed__3; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__5() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____closed__5() { _start: { lean_object* x_1; @@ -7419,22 +7979,22 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Meta_addCustomEliminatorEntry), 2, 0); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__6() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____closed__6() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_SMap_switch___at_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____spec__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_SMap_switch___at_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____spec__1), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__7() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____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_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__4; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__5; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____closed__4; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____closed__5; x_3 = l_Lean_Meta_instInhabitedCustomEliminators___closed__2; -x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__6; +x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____closed__6; x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -7443,11 +8003,11 @@ lean_ctor_set(x_5, 3, x_4); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__7; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____closed__7; x_3 = l_Lean_registerSimpleScopedEnvExtension___rarg(x_2, x_1); return x_3; } @@ -9941,165 +10501,169 @@ return x_110; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_mkCustomEliminator___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_Meta_mkCustomEliminator___lambda__1(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_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, 3); -x_11 = lean_array_get_size(x_10); -x_12 = lean_unsigned_to_nat(0u); -x_13 = lean_unsigned_to_nat(1u); -x_14 = l_Lean_Meta_ElimInfo_targetsPos___default___closed__1; -lean_inc_n(x_11, 2); -x_15 = l_Std_Range_forIn_loop___at_Lean_Meta_mkCustomEliminator___spec__19(x_3, x_10, x_11, x_11, x_12, x_11, x_13, x_14, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_11); -if (lean_obj_tag(x_15) == 0) +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_11 = lean_ctor_get(x_1, 3); +x_12 = lean_array_get_size(x_11); +x_13 = lean_unsigned_to_nat(0u); +x_14 = lean_unsigned_to_nat(1u); +x_15 = l_Lean_Meta_ElimInfo_targetsPos___default___closed__1; +lean_inc_n(x_12, 2); +x_16 = l_Std_Range_forIn_loop___at_Lean_Meta_mkCustomEliminator___spec__19(x_4, x_11, x_12, x_12, x_13, x_12, x_14, x_15, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_12); +if (lean_obj_tag(x_16) == 0) { -uint8_t x_16; -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) +uint8_t x_17; +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 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_17); -lean_ctor_set(x_18, 1, x_2); -lean_ctor_set(x_15, 0, x_18); -return x_15; +lean_object* x_18; lean_object* x_19; +x_18 = lean_ctor_get(x_16, 0); +x_19 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +lean_ctor_set_uint8(x_19, sizeof(void*)*2, x_2); +lean_ctor_set(x_16, 0, x_19); +return x_16; } 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_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_inc(x_19); -lean_dec(x_15); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_19); -lean_ctor_set(x_21, 1, x_2); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_20); -return x_22; +lean_dec(x_16); +x_22 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_3); +lean_ctor_set_uint8(x_22, sizeof(void*)*2, x_2); +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_23; -lean_dec(x_2); -x_23 = !lean_is_exclusive(x_15); -if (x_23 == 0) +uint8_t x_24; +lean_dec(x_3); +x_24 = !lean_is_exclusive(x_16); +if (x_24 == 0) { -return x_15; +return x_16; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_15, 0); -x_25 = lean_ctor_get(x_15, 1); +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_16, 0); +x_26 = lean_ctor_get(x_16, 1); +lean_inc(x_26); lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_15); -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_dec(x_16); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +return x_27; } } } } -LEAN_EXPORT lean_object* l_Lean_Meta_mkCustomEliminator(lean_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_mkCustomEliminator(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_7; lean_object* x_8; -x_7 = lean_box(0); +lean_object* x_8; lean_object* x_9; +x_8 = lean_box(0); +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_getElimInfo(x_1, x_7, x_2, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_8) == 0) +x_9 = l_Lean_Meta_getElimInfo(x_1, x_8, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_9) == 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_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_ctor_get(x_9, 0); lean_inc(x_10); -lean_dec(x_8); -lean_inc(x_1); -x_11 = l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(x_1, x_2, x_3, x_4, x_5, x_10); -if (lean_obj_tag(x_11) == 0) +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +lean_inc(x_1); +x_12 = l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(x_1, x_3, x_4, x_5, x_6, x_11); +if (lean_obj_tag(x_12) == 0) { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_12 = lean_ctor_get(x_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; +x_13 = lean_ctor_get(x_12, 0); lean_inc(x_13); -lean_dec(x_11); -x_14 = l_Lean_ConstantInfo_type(x_12); +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_mkCustomEliminator___lambda__1___boxed), 9, 2); -lean_closure_set(x_15, 0, x_9); -lean_closure_set(x_15, 1, x_1); -x_16 = l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2___rarg(x_14, x_15, x_2, x_3, x_4, x_5, x_13); -return x_16; +x_15 = l_Lean_ConstantInfo_type(x_13); +lean_dec(x_13); +x_16 = lean_box(x_2); +x_17 = lean_alloc_closure((void*)(l_Lean_Meta_mkCustomEliminator___lambda__1___boxed), 10, 3); +lean_closure_set(x_17, 0, x_10); +lean_closure_set(x_17, 1, x_16); +lean_closure_set(x_17, 2, x_1); +x_18 = l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2___rarg(x_15, x_17, x_3, x_4, x_5, x_6, x_14); +return x_18; } else { -uint8_t x_17; -lean_dec(x_9); +uint8_t x_19; +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_dec(x_1); -x_17 = !lean_is_exclusive(x_11); -if (x_17 == 0) +x_19 = !lean_is_exclusive(x_12); +if (x_19 == 0) { -return x_11; +return x_12; } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_11, 0); -x_19 = lean_ctor_get(x_11, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_11); -x_20 = lean_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_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_12, 0); +x_21 = lean_ctor_get(x_12, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_12); +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 { -uint8_t x_21; +uint8_t x_23; +lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_21 = !lean_is_exclusive(x_8); -if (x_21 == 0) +x_23 = !lean_is_exclusive(x_9); +if (x_23 == 0) { -return x_8; +return x_9; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_8, 0); -x_23 = lean_ctor_get(x_8, 1); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_8); -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_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; } } } @@ -10340,22 +10904,34 @@ lean_dec(x_1); return x_14; } } -LEAN_EXPORT lean_object* l_Lean_Meta_mkCustomEliminator___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_Meta_mkCustomEliminator___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_10; -x_10 = l_Lean_Meta_mkCustomEliminator___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +uint8_t x_11; lean_object* x_12; +x_11 = lean_unbox(x_2); +lean_dec(x_2); +x_12 = l_Lean_Meta_mkCustomEliminator___lambda__1(x_1, x_11, 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); lean_dec(x_1); -return x_10; +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_mkCustomEliminator___boxed(lean_object* x_1, 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_Meta_mkCustomEliminator(x_1, x_8, x_3, x_4, x_5, x_6, x_7); +return x_9; } } static lean_object* _init_l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addCustomEliminator___spec__1___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_CustomEliminators_map___default___closed__4; +x_1 = l_Lean_Meta_CustomEliminators_map___default___closed__9; x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10378,7 +10954,7 @@ static lean_object* _init_l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addCustom _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_CustomEliminators_map___default___closed__4; +x_1 = l_Lean_Meta_CustomEliminators_map___default___closed__9; x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10398,7 +10974,7 @@ static lean_object* _init_l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addCustom _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_CustomEliminators_map___default___closed__4; +x_1 = l_Lean_Meta_CustomEliminators_map___default___closed__9; x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10470,7 +11046,7 @@ static lean_object* _init_l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addCustom _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_CustomEliminators_map___default___closed__4; +x_1 = l_Lean_Meta_CustomEliminators_map___default___closed__9; x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10504,7 +11080,7 @@ static lean_object* _init_l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addCustom _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_CustomEliminators_map___default___closed__4; +x_1 = l_Lean_Meta_CustomEliminators_map___default___closed__9; x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11116,54 +11692,54 @@ x_1 = l_Lean_Meta_customEliminatorExt; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_addCustomEliminator(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_Meta_addCustomEliminator(lean_object* x_1, uint8_t 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) { _start: { -lean_object* x_8; +lean_object* x_9; +lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_3); -x_8 = l_Lean_Meta_mkCustomEliminator(x_1, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_8) == 0) +x_9 = l_Lean_Meta_mkCustomEliminator(x_1, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_9) == 0) { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_9 = lean_ctor_get(x_8, 0); -lean_inc(x_9); -x_10 = lean_ctor_get(x_8, 1); +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_10 = lean_ctor_get(x_9, 0); lean_inc(x_10); -lean_dec(x_8); -x_11 = l_Lean_Meta_addCustomEliminator___closed__1; -x_12 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addCustomEliminator___spec__1(x_11, x_9, x_2, x_3, x_4, x_5, x_6, x_10); -lean_dec(x_6); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +x_12 = l_Lean_Meta_addCustomEliminator___closed__1; +x_13 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addCustomEliminator___spec__1(x_12, x_10, x_2, x_4, x_5, x_6, x_7, x_11); +lean_dec(x_7); +lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -return x_12; +return x_13; } else { -uint8_t x_13; +uint8_t x_14; +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -x_13 = !lean_is_exclusive(x_8); -if (x_13 == 0) +x_14 = !lean_is_exclusive(x_9); +if (x_14 == 0) { -return x_8; +return x_9; } 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_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_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; +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; } } } @@ -11181,17 +11757,19 @@ lean_dec(x_4); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Meta_addCustomEliminator___boxed(lean_object* x_1, lean_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_addCustomEliminator___boxed(lean_object* x_1, lean_object* x_2, 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_8; lean_object* x_9; -x_8 = lean_unbox(x_2); +uint8_t x_9; uint8_t x_10; lean_object* x_11; +x_9 = lean_unbox(x_2); lean_dec(x_2); -x_9 = l_Lean_Meta_addCustomEliminator(x_1, x_8, x_3, x_4, x_5, x_6, x_7); -return x_9; +x_10 = lean_unbox(x_3); +lean_dec(x_3); +x_11 = l_Lean_Meta_addCustomEliminator(x_1, x_9, x_10, x_4, x_5, x_6, x_7, x_8); +return x_11; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__1() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__1() { _start: { uint8_t x_1; uint8_t x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; @@ -11215,11 +11793,11 @@ lean_ctor_set_uint8(x_5, 11, x_4); return x_5; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__2() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_CustomEliminators_map___default___closed__4; +x_1 = l_Lean_Meta_CustomEliminators_map___default___closed__9; x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11227,7 +11805,7 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__3() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__3() { _start: { lean_object* x_1; lean_object* x_2; @@ -11236,23 +11814,23 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__4() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__3; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__3; 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_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__5() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__5() { _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_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__4; -x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__3; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__4; +x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__3; 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); @@ -11263,25 +11841,25 @@ lean_ctor_set_usize(x_5, 4, x_1); return x_5; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__6() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__2; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__5; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__2; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__5; x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__7() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____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; x_1 = lean_box(0); -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__1; -x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__6; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__1; +x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__6; x_4 = l_Lean_Meta_ElimInfo_targetsPos___default___closed__1; x_5 = lean_unsigned_to_nat(0u); x_6 = lean_alloc_ctor(0, 6, 0); @@ -11294,11 +11872,11 @@ lean_ctor_set(x_6, 5, x_1); return x_6; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__8() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_CustomEliminators_map___default___closed__4; +x_1 = l_Lean_Meta_CustomEliminators_map___default___closed__9; x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11306,11 +11884,11 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__9() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_CustomEliminators_map___default___closed__4; +x_1 = l_Lean_Meta_CustomEliminators_map___default___closed__9; x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11318,13 +11896,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__10() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____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 = lean_unsigned_to_nat(0u); -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__8; -x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__9; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__8; +x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__9; x_4 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addCustomEliminator___spec__1___closed__1; x_5 = lean_alloc_ctor(0, 9, 0); lean_ctor_set(x_5, 0, x_1); @@ -11339,14 +11917,14 @@ lean_ctor_set(x_5, 8, x_3); return x_5; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__11() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___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 = lean_box(0); -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__10; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__10; x_3 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addCustomEliminator___spec__1___closed__17; -x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__5; +x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__5; x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_2); lean_ctor_set(x_5, 1, x_3); @@ -11355,77 +11933,78 @@ lean_ctor_set(x_5, 3, x_4); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_7 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__11; +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; +x_7 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__11; x_8 = lean_st_mk_ref(x_7, x_6); 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_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__7; +x_11 = 1; +x_12 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__7; lean_inc(x_9); -x_12 = l_Lean_Meta_addCustomEliminator(x_1, x_3, x_11, x_9, x_4, x_5, x_10); -if (lean_obj_tag(x_12) == 0) +x_13 = l_Lean_Meta_addCustomEliminator(x_1, x_3, x_11, x_12, x_9, x_4, x_5, x_10); +if (lean_obj_tag(x_13) == 0) { -lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_13 = lean_ctor_get(x_12, 1); -lean_inc(x_13); -lean_dec(x_12); -x_14 = lean_st_ref_get(x_9, x_13); +lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_14 = lean_ctor_get(x_13, 1); +lean_inc(x_14); +lean_dec(x_13); +x_15 = lean_st_ref_get(x_9, x_14); lean_dec(x_9); -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) { -lean_object* x_16; lean_object* x_17; -x_16 = lean_ctor_get(x_14, 0); -lean_dec(x_16); -x_17 = lean_box(0); -lean_ctor_set(x_14, 0, x_17); -return x_14; +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_15, 0); +lean_dec(x_17); +x_18 = lean_box(0); +lean_ctor_set(x_15, 0, x_18); +return x_15; } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_14, 1); -lean_inc(x_18); -lean_dec(x_14); -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_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_15, 1); +lean_inc(x_19); +lean_dec(x_15); +x_20 = lean_box(0); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_19); +return x_21; } } else { -uint8_t x_21; +uint8_t x_22; lean_dec(x_9); -x_21 = !lean_is_exclusive(x_12); -if (x_21 == 0) +x_22 = !lean_is_exclusive(x_13); +if (x_22 == 0) { -return x_12; +return x_13; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_12, 0); -x_23 = lean_ctor_get(x_12, 1); +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_13, 0); +x_24 = lean_ctor_get(x_13, 1); +lean_inc(x_24); lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_12); -x_24 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_24, 0, x_22); -lean_ctor_set(x_24, 1, x_23); -return x_24; +lean_dec(x_13); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; } } } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__2___closed__1() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__2___closed__1() { _start: { lean_object* x_1; @@ -11433,45 +12012,45 @@ x_1 = lean_mk_string_from_bytes("attribute cannot be erased", 26); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__2___closed__2() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__2___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__2___closed__1; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__2___closed__1; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; -x_5 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__2___closed__2; +x_5 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__2___closed__2; x_6 = l_Lean_throwError___at_Lean_AttributeImpl_erase___default___spec__1(x_5, x_2, x_3, x_4); return x_6; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__1() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____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_Tactic_ElimInfo___hyg_2298____closed__1; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____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_Tactic_ElimInfo___hyg_2946____closed__2() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__1; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__2; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__1; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____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_Tactic_ElimInfo___hyg_2946____closed__3() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__3() { _start: { lean_object* x_1; @@ -11479,17 +12058,17 @@ x_1 = lean_mk_string_from_bytes("initFn", 6); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__4() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__2; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__3; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__2; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____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_Tactic_ElimInfo___hyg_2946____closed__5() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__5() { _start: { lean_object* x_1; @@ -11497,37 +12076,37 @@ x_1 = lean_mk_string_from_bytes("_@", 2); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__6() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__4; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__5; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__4; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____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_Tactic_ElimInfo___hyg_2946____closed__7() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__6; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__1; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__6; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____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_Tactic_ElimInfo___hyg_2946____closed__8() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__7; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__2; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__7; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____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_Tactic_ElimInfo___hyg_2946____closed__9() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__9() { _start: { lean_object* x_1; @@ -11535,17 +12114,17 @@ x_1 = lean_mk_string_from_bytes("Tactic", 6); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__10() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__8; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__9; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__8; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____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_Tactic_ElimInfo___hyg_2946____closed__11() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__11() { _start: { lean_object* x_1; @@ -11553,17 +12132,17 @@ x_1 = lean_mk_string_from_bytes("ElimInfo", 8); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__12() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__10; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__11; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__10; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____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_Tactic_ElimInfo___hyg_2946____closed__13() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__13() { _start: { lean_object* x_1; @@ -11571,59 +12150,59 @@ x_1 = lean_mk_string_from_bytes("_hyg", 4); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__14() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__12; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__13; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__12; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__13; 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_Tactic_ElimInfo___hyg_2946____closed__15() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__14; -x_2 = lean_unsigned_to_nat(2946u); +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__14; +x_2 = lean_unsigned_to_nat(2991u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__16() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__16() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("eliminator", 10); +x_1 = lean_mk_string_from_bytes("induction_eliminator", 20); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__17() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__17() { _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_Tactic_ElimInfo___hyg_2946____closed__16; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__16; 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_Tactic_ElimInfo___hyg_2946____closed__18() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__18() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("custom eliminator for `cases` and `induction` tactics", 53); +x_1 = lean_mk_string_from_bytes("custom `rec`-like eliminator for the `induction` tactic", 55); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__19() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____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_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__15; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__17; -x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__18; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__15; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__17; +x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__18; x_4 = 0; x_5 = lean_alloc_ctor(0, 3, 1); lean_ctor_set(x_5, 0, x_1); @@ -11633,29 +12212,29 @@ lean_ctor_set_uint8(x_5, sizeof(void*)*3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__20() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__20() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___boxed), 6, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___boxed), 6, 0); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__21() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__21() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__2___boxed), 4, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__2___boxed), 4, 0); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__22() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__22() { _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_Tactic_ElimInfo___hyg_2946____closed__19; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__20; -x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__21; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__19; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__20; +x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__21; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -11663,97 +12242,262 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__22; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__22; x_3 = l_Lean_registerBuiltinAttribute(x_2, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____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_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { uint8_t x_7; lean_object* x_8; x_7 = lean_unbox(x_3); lean_dec(x_3); -x_8 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1(x_1, x_2, x_7, x_4, x_5, x_6); +x_8 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1(x_1, x_2, x_7, x_4, x_5, x_6); lean_dec(x_2); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__2(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__2(x_1, x_2, x_3, x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Meta_getCustomEliminators___rarg(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____lambda__1(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_3; uint8_t x_4; -x_3 = lean_st_ref_get(x_1, x_2); -x_4 = !lean_is_exclusive(x_3); -if (x_4 == 0) +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; +x_7 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__11; +x_8 = lean_st_mk_ref(x_7, x_6); +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 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__7; +lean_inc(x_9); +x_13 = l_Lean_Meta_addCustomEliminator(x_1, x_3, x_11, x_12, x_9, x_4, x_5, x_10); +if (lean_obj_tag(x_13) == 0) { -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_3, 0); -x_6 = lean_ctor_get(x_5, 0); -lean_inc(x_6); -lean_dec(x_5); -x_7 = l_Lean_Meta_instInhabitedCustomEliminators; -x_8 = l_Lean_Meta_addCustomEliminator___closed__1; -x_9 = l_Lean_ScopedEnvExtension_getState___rarg(x_7, x_8, x_6); -lean_dec(x_6); -lean_ctor_set(x_3, 0, x_9); -return x_3; +lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_14 = lean_ctor_get(x_13, 1); +lean_inc(x_14); +lean_dec(x_13); +x_15 = lean_st_ref_get(x_9, x_14); +lean_dec(x_9); +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_15, 0); +lean_dec(x_17); +x_18 = lean_box(0); +lean_ctor_set(x_15, 0, x_18); +return x_15; } 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; lean_object* x_16; -x_10 = lean_ctor_get(x_3, 0); -x_11 = lean_ctor_get(x_3, 1); -lean_inc(x_11); -lean_inc(x_10); -lean_dec(x_3); -x_12 = lean_ctor_get(x_10, 0); -lean_inc(x_12); -lean_dec(x_10); -x_13 = l_Lean_Meta_instInhabitedCustomEliminators; -x_14 = l_Lean_Meta_addCustomEliminator___closed__1; -x_15 = l_Lean_ScopedEnvExtension_getState___rarg(x_13, x_14, x_12); -lean_dec(x_12); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_11); -return x_16; +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_15, 1); +lean_inc(x_19); +lean_dec(x_15); +x_20 = lean_box(0); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_19); +return x_21; } } +else +{ +uint8_t x_22; +lean_dec(x_9); +x_22 = !lean_is_exclusive(x_13); +if (x_22 == 0) +{ +return x_13; } -LEAN_EXPORT lean_object* l_Lean_Meta_getCustomEliminators(lean_object* x_1) { -_start: +else { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Meta_getCustomEliminators___rarg___boxed), 2, 0); -return x_2; +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_13, 0); +x_24 = lean_ctor_get(x_13, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_13); +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_Meta_getCustomEliminators___rarg___boxed(lean_object* x_1, lean_object* x_2) { +} +} +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____closed__1() { _start: { -lean_object* x_3; -x_3 = l_Lean_Meta_getCustomEliminators___rarg(x_1, x_2); -lean_dec(x_1); +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__14; +x_2 = lean_unsigned_to_nat(3070u); +x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_getCustomEliminators___boxed(lean_object* x_1) { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("cases_eliminator", 16); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____closed__3() { +_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_Tactic_ElimInfo___hyg_3070____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_Tactic_ElimInfo___hyg_3070____closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("custom `casesOn`-like eliminator for the `cases` tactic", 55); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; uint8_t x_4; lean_object* x_5; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____closed__1; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____closed__3; +x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____closed__4; +x_4 = 0; +x_5 = lean_alloc_ctor(0, 3, 1); +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_uint8(x_5, sizeof(void*)*3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____lambda__1___boxed), 6, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____closed__7() { +_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_Tactic_ElimInfo___hyg_3070____closed__5; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____closed__6; +x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__21; +x_4 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070_(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____closed__7; +x_3 = l_Lean_registerBuiltinAttribute(x_2, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; lean_object* x_8; +x_7 = lean_unbox(x_3); +lean_dec(x_3); +x_8 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____lambda__1(x_1, x_2, x_7, x_4, x_5, x_6); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_getCustomEliminators___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; uint8_t x_4; +x_3 = lean_st_ref_get(x_1, x_2); +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +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_3, 0); +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +lean_dec(x_5); +x_7 = l_Lean_Meta_instInhabitedCustomEliminators; +x_8 = l_Lean_Meta_addCustomEliminator___closed__1; +x_9 = l_Lean_ScopedEnvExtension_getState___rarg(x_7, x_8, x_6); +lean_dec(x_6); +lean_ctor_set(x_3, 0, x_9); +return x_3; +} +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; lean_object* x_16; +x_10 = lean_ctor_get(x_3, 0); +x_11 = lean_ctor_get(x_3, 1); +lean_inc(x_11); +lean_inc(x_10); +lean_dec(x_3); +x_12 = lean_ctor_get(x_10, 0); +lean_inc(x_12); +lean_dec(x_10); +x_13 = l_Lean_Meta_instInhabitedCustomEliminators; +x_14 = l_Lean_Meta_addCustomEliminator___closed__1; +x_15 = l_Lean_ScopedEnvExtension_getState___rarg(x_13, x_14, x_12); +lean_dec(x_12); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_11); +return x_16; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_getCustomEliminators(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_getCustomEliminators___rarg___boxed), 2, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_getCustomEliminators___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Meta_getCustomEliminators___rarg(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_getCustomEliminators___boxed(lean_object* x_1) { _start: { lean_object* x_2; @@ -11956,44 +12700,43 @@ goto _start; } } } -LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Meta_getCustomEliminator_x3f___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Meta_getCustomEliminator_x3f___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_9; uint8_t x_10; -x_9 = lean_array_get_size(x_6); -x_10 = lean_nat_dec_lt(x_8, x_9); -lean_dec(x_9); -if (x_10 == 0) +lean_object* x_7; uint8_t x_8; +x_7 = lean_array_get_size(x_4); +x_8 = lean_nat_dec_lt(x_6, x_7); +lean_dec(x_7); +if (x_8 == 0) { -uint8_t x_11; -lean_dec(x_8); -x_11 = 1; -return x_11; +uint8_t x_9; +lean_dec(x_6); +x_9 = 1; +return x_9; } else { -lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_12 = lean_array_fget(x_6, x_8); -x_13 = lean_array_fget(x_7, x_8); -x_14 = lean_name_eq(x_12, x_13); -lean_dec(x_13); -lean_dec(x_12); -if (x_14 == 0) +lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_10 = lean_array_fget(x_4, x_6); +x_11 = lean_array_fget(x_5, x_6); +x_12 = lean_name_eq(x_10, x_11); +lean_dec(x_11); +lean_dec(x_10); +if (x_12 == 0) { -uint8_t x_15; -lean_dec(x_8); -x_15 = 0; -return x_15; +uint8_t x_13; +lean_dec(x_6); +x_13 = 0; +return x_13; } else { -lean_object* x_16; lean_object* x_17; -x_16 = lean_unsigned_to_nat(1u); -x_17 = lean_nat_add(x_8, x_16); -lean_dec(x_8); -x_4 = lean_box(0); -x_5 = lean_box(0); -x_8 = x_17; +lean_object* x_14; lean_object* x_15; +x_14 = lean_unsigned_to_nat(1u); +x_15 = lean_nat_add(x_6, x_14); +lean_dec(x_6); +x_3 = lean_box(0); +x_6 = x_15; goto _start; } } @@ -12015,48 +12758,114 @@ return x_8; } else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; uint8_t x_32; x_9 = lean_array_fget(x_1, x_4); -x_10 = lean_array_get_size(x_5); -x_11 = lean_array_get_size(x_9); -x_12 = lean_nat_dec_eq(x_10, x_11); -lean_dec(x_11); -lean_dec(x_10); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; +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_13 = lean_unsigned_to_nat(1u); -x_14 = lean_nat_add(x_4, x_13); +x_32 = lean_unbox(x_10); +if (x_32 == 0) +{ +uint8_t x_33; +x_33 = lean_unbox(x_12); +lean_dec(x_12); +if (x_33 == 0) +{ +uint8_t x_34; +x_34 = 1; +x_14 = x_34; +goto block_31; +} +else +{ +uint8_t x_35; +x_35 = 0; +x_14 = x_35; +goto block_31; +} +} +else +{ +uint8_t x_36; +x_36 = lean_unbox(x_12); +lean_dec(x_12); +if (x_36 == 0) +{ +uint8_t x_37; +x_37 = 0; +x_14 = x_37; +goto block_31; +} +else +{ +uint8_t x_38; +x_38 = 1; +x_14 = x_38; +goto block_31; +} +} +block_31: +{ +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +lean_dec(x_13); +x_15 = lean_unsigned_to_nat(1u); +x_16 = lean_nat_add(x_4, x_15); lean_dec(x_4); x_3 = lean_box(0); -x_4 = x_14; +x_4 = x_16; goto _start; } else { -lean_object* x_16; uint8_t x_17; -x_16 = lean_unsigned_to_nat(0u); -x_17 = l_Array_isEqvAux___at_Lean_Meta_getCustomEliminator_x3f___spec__7(x_1, x_4, x_5, lean_box(0), lean_box(0), x_5, x_9, x_16); -lean_dec(x_9); -if (x_17 == 0) +lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_18 = lean_array_get_size(x_11); +x_19 = lean_array_get_size(x_13); +x_20 = lean_nat_dec_eq(x_18, x_19); +lean_dec(x_19); +lean_dec(x_18); +if (x_20 == 0) { -lean_object* x_18; lean_object* x_19; -x_18 = lean_unsigned_to_nat(1u); -x_19 = lean_nat_add(x_4, x_18); +lean_object* x_21; lean_object* x_22; +lean_dec(x_13); +x_21 = lean_unsigned_to_nat(1u); +x_22 = lean_nat_add(x_4, x_21); lean_dec(x_4); x_3 = lean_box(0); -x_4 = x_19; +x_4 = x_22; goto _start; } else { -lean_object* x_21; lean_object* x_22; -x_21 = lean_array_fget(x_2, x_4); +lean_object* x_24; uint8_t x_25; +x_24 = lean_unsigned_to_nat(0u); +x_25 = l_Array_isEqvAux___at_Lean_Meta_getCustomEliminator_x3f___spec__7(x_11, x_13, lean_box(0), x_11, x_13, x_24); +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_22 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_22, 0, x_21); -return x_22; +x_3 = lean_box(0); +x_4 = x_27; +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; +} +} } } } @@ -12082,126 +12891,246 @@ lean_dec(x_4); switch (lean_obj_tag(x_10)) { case 0: { -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_16; uint8_t x_17; uint8_t x_28; 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_get_size(x_3); -x_14 = lean_array_get_size(x_11); -x_15 = lean_nat_dec_eq(x_13, x_14); -lean_dec(x_14); -lean_dec(x_13); -if (x_15 == 0) +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_28 = lean_unbox(x_13); +if (x_28 == 0) { -lean_object* x_16; +uint8_t x_29; +x_29 = lean_unbox(x_15); +lean_dec(x_15); +if (x_29 == 0) +{ +uint8_t x_30; +x_30 = 1; +x_17 = x_30; +goto block_27; +} +else +{ +uint8_t x_31; +x_31 = 0; +x_17 = x_31; +goto block_27; +} +} +else +{ +uint8_t x_32; +x_32 = lean_unbox(x_15); +lean_dec(x_15); +if (x_32 == 0) +{ +uint8_t x_33; +x_33 = 0; +x_17 = x_33; +goto block_27; +} +else +{ +uint8_t x_34; +x_34 = 1; +x_17 = x_34; +goto block_27; +} +} +block_27: +{ +if (x_17 == 0) +{ +lean_object* x_18; +lean_dec(x_16); lean_dec(x_12); -lean_dec(x_11); -x_16 = lean_box(0); -return x_16; +x_18 = lean_box(0); +return x_18; } else { -lean_object* x_17; uint8_t x_18; -x_17 = lean_unsigned_to_nat(0u); -x_18 = l_Array_isEqvAux___at_Lean_Meta_getCustomEliminator_x3f___spec__5(x_3, x_11, lean_box(0), x_3, x_11, x_17); -lean_dec(x_11); -if (x_18 == 0) +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = lean_array_get_size(x_14); +x_20 = lean_array_get_size(x_16); +x_21 = lean_nat_dec_eq(x_19, x_20); +lean_dec(x_20); +lean_dec(x_19); +if (x_21 == 0) { -lean_object* x_19; +lean_object* x_22; +lean_dec(x_16); lean_dec(x_12); -x_19 = lean_box(0); -return x_19; +x_22 = lean_box(0); +return x_22; } else { -lean_object* x_20; -x_20 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_20, 0, x_12); -return x_20; +lean_object* x_23; uint8_t x_24; +x_23 = lean_unsigned_to_nat(0u); +x_24 = l_Array_isEqvAux___at_Lean_Meta_getCustomEliminator_x3f___spec__5(x_14, x_16, lean_box(0), x_14, x_16, x_23); +lean_dec(x_16); +if (x_24 == 0) +{ +lean_object* x_25; +lean_dec(x_12); +x_25 = lean_box(0); +return x_25; +} +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_21; size_t x_22; -x_21 = lean_ctor_get(x_10, 0); -lean_inc(x_21); +lean_object* x_35; size_t x_36; +x_35 = lean_ctor_get(x_10, 0); +lean_inc(x_35); lean_dec(x_10); -x_22 = lean_usize_shift_right(x_2, x_5); -x_1 = x_21; -x_2 = x_22; +x_36 = lean_usize_shift_right(x_2, x_5); +x_1 = x_35; +x_2 = x_36; goto _start; } default: { -lean_object* x_24; -x_24 = lean_box(0); -return x_24; +lean_object* x_38; +x_38 = lean_box(0); +return x_38; } } } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_25 = lean_ctor_get(x_1, 0); -lean_inc(x_25); -x_26 = lean_ctor_get(x_1, 1); -lean_inc(x_26); +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_39 = lean_ctor_get(x_1, 0); +lean_inc(x_39); +x_40 = lean_ctor_get(x_1, 1); +lean_inc(x_40); lean_dec(x_1); -x_27 = lean_unsigned_to_nat(0u); -x_28 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_getCustomEliminator_x3f___spec__6(x_25, x_26, lean_box(0), x_27, x_3); -lean_dec(x_26); -lean_dec(x_25); -return x_28; +x_41 = lean_unsigned_to_nat(0u); +x_42 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_getCustomEliminator_x3f___spec__6(x_39, x_40, lean_box(0), x_41, x_3); +lean_dec(x_40); +lean_dec(x_39); +return x_42; } } } LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_getCustomEliminator_x3f___spec__3(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_3; uint64_t x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +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; uint8_t x_10; x_3 = lean_ctor_get(x_1, 0); lean_inc(x_3); lean_dec(x_1); -x_4 = 7; -x_5 = lean_array_get_size(x_2); -x_6 = lean_unsigned_to_nat(0u); -x_7 = lean_nat_dec_lt(x_6, x_5); -if (x_7 == 0) +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_5); +x_8 = lean_unsigned_to_nat(0u); +x_9 = lean_nat_dec_lt(x_8, x_7); +x_10 = lean_unbox(x_4); +lean_dec(x_4); +if (x_10 == 0) +{ +if (x_9 == 0) { -size_t x_8; lean_object* x_9; +size_t x_11; lean_object* x_12; +lean_dec(x_7); lean_dec(x_5); -x_8 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__1; -x_9 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_getCustomEliminator_x3f___spec__4(x_3, x_8, x_2); +x_11 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__2; +x_12 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_getCustomEliminator_x3f___spec__4(x_3, x_11, x_2); lean_dec(x_2); -return x_9; +return x_12; } else { -uint8_t x_10; -x_10 = lean_nat_dec_le(x_5, x_5); -if (x_10 == 0) +uint64_t x_13; uint8_t x_14; +x_13 = 13; +x_14 = lean_nat_dec_le(x_7, x_7); +if (x_14 == 0) { -size_t x_11; lean_object* x_12; +size_t x_15; lean_object* x_16; +lean_dec(x_7); lean_dec(x_5); -x_11 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__1; -x_12 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_getCustomEliminator_x3f___spec__4(x_3, x_11, x_2); +x_15 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__2; +x_16 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_getCustomEliminator_x3f___spec__4(x_3, x_15, x_2); lean_dec(x_2); -return x_12; +return x_16; } else { -size_t x_13; size_t x_14; uint64_t x_15; size_t x_16; lean_object* x_17; -x_13 = 0; -x_14 = lean_usize_of_nat(x_5); +size_t x_17; size_t x_18; uint64_t x_19; uint64_t x_20; size_t x_21; lean_object* x_22; +x_17 = 0; +x_18 = lean_usize_of_nat(x_7); +lean_dec(x_7); +x_19 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_addCustomEliminatorEntry___spec__5(x_5, x_17, x_18, x_6); lean_dec(x_5); -x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_addCustomEliminatorEntry___spec__5(x_2, x_13, x_14, x_4); -x_16 = lean_uint64_to_usize(x_15); -x_17 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_getCustomEliminator_x3f___spec__4(x_3, x_16, x_2); +x_20 = lean_uint64_mix_hash(x_13, x_19); +x_21 = lean_uint64_to_usize(x_20); +x_22 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_getCustomEliminator_x3f___spec__4(x_3, x_21, x_2); lean_dec(x_2); -return x_17; +return x_22; +} +} +} +else +{ +if (x_9 == 0) +{ +size_t x_23; lean_object* x_24; +lean_dec(x_7); +lean_dec(x_5); +x_23 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__4; +x_24 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_getCustomEliminator_x3f___spec__4(x_3, x_23, x_2); +lean_dec(x_2); +return x_24; +} +else +{ +uint64_t x_25; uint8_t x_26; +x_25 = 11; +x_26 = lean_nat_dec_le(x_7, x_7); +if (x_26 == 0) +{ +size_t x_27; lean_object* x_28; +lean_dec(x_7); +lean_dec(x_5); +x_27 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__4; +x_28 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_getCustomEliminator_x3f___spec__4(x_3, x_27, x_2); +lean_dec(x_2); +return x_28; +} +else +{ +size_t x_29; size_t x_30; uint64_t x_31; uint64_t x_32; size_t x_33; lean_object* x_34; +x_29 = 0; +x_30 = lean_usize_of_nat(x_7); +lean_dec(x_7); +x_31 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_addCustomEliminatorEntry___spec__5(x_5, x_29, x_30, x_6); +lean_dec(x_5); +x_32 = lean_uint64_mix_hash(x_25, x_31); +x_33 = lean_uint64_to_usize(x_32); +x_34 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_getCustomEliminator_x3f___spec__4(x_3, x_33, x_2); +lean_dec(x_2); +return x_34; +} } } } @@ -12248,48 +13177,103 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_Meta_getCustomEliminator_x3f___spec__9(lean_object* x_1, lean_object* x_2) { -_start: +LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_Meta_getCustomEliminator_x3f___spec__9(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; +x_3 = lean_box(0); +return x_3; +} +else +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; uint8_t x_22; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 1); +x_6 = lean_ctor_get(x_2, 2); +x_7 = lean_ctor_get(x_4, 0); +x_8 = lean_ctor_get(x_4, 1); +x_9 = lean_ctor_get(x_1, 0); +x_10 = lean_ctor_get(x_1, 1); +x_22 = lean_unbox(x_7); +if (x_22 == 0) +{ +uint8_t x_23; +x_23 = lean_unbox(x_9); +if (x_23 == 0) +{ +uint8_t x_24; +x_24 = 1; +x_11 = x_24; +goto block_21; +} +else +{ +uint8_t x_25; +x_25 = 0; +x_11 = x_25; +goto block_21; +} +} +else +{ +uint8_t x_26; +x_26 = lean_unbox(x_9); +if (x_26 == 0) +{ +uint8_t x_27; +x_27 = 0; +x_11 = x_27; +goto block_21; +} +else +{ +uint8_t x_28; +x_28 = 1; +x_11 = x_28; +goto block_21; +} +} +block_21: { -if (lean_obj_tag(x_2) == 0) +if (x_11 == 0) { -lean_object* x_3; -x_3 = lean_box(0); -return x_3; +x_2 = x_6; +goto _start; } else { -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; -x_4 = lean_ctor_get(x_2, 0); -x_5 = lean_ctor_get(x_2, 1); -x_6 = lean_ctor_get(x_2, 2); -x_7 = lean_array_get_size(x_4); -x_8 = lean_array_get_size(x_1); -x_9 = lean_nat_dec_eq(x_7, x_8); -lean_dec(x_8); -lean_dec(x_7); -if (x_9 == 0) +lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_13 = lean_array_get_size(x_8); +x_14 = lean_array_get_size(x_10); +x_15 = lean_nat_dec_eq(x_13, x_14); +lean_dec(x_14); +lean_dec(x_13); +if (x_15 == 0) { x_2 = x_6; goto _start; } else { -lean_object* x_11; uint8_t x_12; -x_11 = lean_unsigned_to_nat(0u); -x_12 = l_Array_isEqvAux___at_Lean_Meta_getCustomEliminator_x3f___spec__10(x_1, x_4, lean_box(0), x_4, x_1, x_11); -if (x_12 == 0) +lean_object* x_17; uint8_t x_18; +x_17 = lean_unsigned_to_nat(0u); +x_18 = l_Array_isEqvAux___at_Lean_Meta_getCustomEliminator_x3f___spec__10(x_8, x_10, lean_box(0), x_8, x_10, x_17); +if (x_18 == 0) { x_2 = x_6; goto _start; } else { -lean_object* x_14; +lean_object* x_20; lean_inc(x_5); -x_14 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_14, 0, x_5); -return x_14; +x_20 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_20, 0, x_5); +return x_20; +} +} } } } @@ -12298,36 +13282,30 @@ return x_14; LEAN_EXPORT lean_object* l_Lean_HashMapImp_find_x3f___at_Lean_Meta_getCustomEliminator_x3f___spec__8(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_3; lean_object* x_4; uint64_t x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint64_t x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; uint8_t x_11; x_3 = lean_ctor_get(x_1, 1); lean_inc(x_3); lean_dec(x_1); x_4 = lean_array_get_size(x_3); -x_5 = 7; -x_6 = lean_array_get_size(x_2); -x_7 = lean_unsigned_to_nat(0u); -x_8 = lean_nat_dec_lt(x_7, x_6); -if (x_8 == 0) -{ -size_t x_9; lean_object* x_10; lean_object* x_11; -lean_dec(x_6); -x_9 = lean_hashmap_mk_idx(x_4, x_5); -x_10 = lean_array_uget(x_3, x_9); -lean_dec(x_3); -x_11 = l_Lean_AssocList_find_x3f___at_Lean_Meta_getCustomEliminator_x3f___spec__9(x_2, x_10); -lean_dec(x_10); -lean_dec(x_2); -return x_11; -} -else +x_5 = lean_ctor_get(x_2, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_2, 1); +lean_inc(x_6); +x_7 = 7; +x_8 = lean_array_get_size(x_6); +x_9 = lean_unsigned_to_nat(0u); +x_10 = lean_nat_dec_lt(x_9, x_8); +x_11 = lean_unbox(x_5); +lean_dec(x_5); +if (x_11 == 0) { -uint8_t x_12; -x_12 = lean_nat_dec_le(x_6, x_6); -if (x_12 == 0) +if (x_10 == 0) { -size_t x_13; lean_object* x_14; lean_object* x_15; +uint64_t x_12; size_t x_13; lean_object* x_14; lean_object* x_15; +lean_dec(x_8); lean_dec(x_6); -x_13 = lean_hashmap_mk_idx(x_4, x_5); +x_12 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__1; +x_13 = lean_hashmap_mk_idx(x_4, x_12); x_14 = lean_array_uget(x_3, x_13); lean_dec(x_3); x_15 = l_Lean_AssocList_find_x3f___at_Lean_Meta_getCustomEliminator_x3f___spec__9(x_2, x_14); @@ -12337,11 +13315,15 @@ return x_15; } else { -size_t x_16; size_t x_17; uint64_t x_18; size_t x_19; lean_object* x_20; lean_object* x_21; -x_16 = 0; -x_17 = lean_usize_of_nat(x_6); +uint64_t x_16; uint8_t x_17; +x_16 = 13; +x_17 = lean_nat_dec_le(x_8, x_8); +if (x_17 == 0) +{ +uint64_t x_18; size_t x_19; lean_object* x_20; lean_object* x_21; +lean_dec(x_8); lean_dec(x_6); -x_18 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_addCustomEliminatorEntry___spec__5(x_2, x_16, x_17, x_5); +x_18 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__1; x_19 = lean_hashmap_mk_idx(x_4, x_18); x_20 = lean_array_uget(x_3, x_19); lean_dec(x_3); @@ -12350,6 +13332,78 @@ lean_dec(x_20); lean_dec(x_2); return x_21; } +else +{ +size_t x_22; size_t x_23; uint64_t x_24; uint64_t x_25; size_t x_26; lean_object* x_27; lean_object* x_28; +x_22 = 0; +x_23 = lean_usize_of_nat(x_8); +lean_dec(x_8); +x_24 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_addCustomEliminatorEntry___spec__5(x_6, x_22, x_23, x_7); +lean_dec(x_6); +x_25 = lean_uint64_mix_hash(x_16, x_24); +x_26 = lean_hashmap_mk_idx(x_4, x_25); +x_27 = lean_array_uget(x_3, x_26); +lean_dec(x_3); +x_28 = l_Lean_AssocList_find_x3f___at_Lean_Meta_getCustomEliminator_x3f___spec__9(x_2, x_27); +lean_dec(x_27); +lean_dec(x_2); +return x_28; +} +} +} +else +{ +if (x_10 == 0) +{ +uint64_t x_29; size_t x_30; lean_object* x_31; lean_object* x_32; +lean_dec(x_8); +lean_dec(x_6); +x_29 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__3; +x_30 = lean_hashmap_mk_idx(x_4, x_29); +x_31 = lean_array_uget(x_3, x_30); +lean_dec(x_3); +x_32 = l_Lean_AssocList_find_x3f___at_Lean_Meta_getCustomEliminator_x3f___spec__9(x_2, x_31); +lean_dec(x_31); +lean_dec(x_2); +return x_32; +} +else +{ +uint64_t x_33; uint8_t x_34; +x_33 = 11; +x_34 = lean_nat_dec_le(x_8, x_8); +if (x_34 == 0) +{ +uint64_t x_35; size_t x_36; lean_object* x_37; lean_object* x_38; +lean_dec(x_8); +lean_dec(x_6); +x_35 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__3; +x_36 = lean_hashmap_mk_idx(x_4, x_35); +x_37 = lean_array_uget(x_3, x_36); +lean_dec(x_3); +x_38 = l_Lean_AssocList_find_x3f___at_Lean_Meta_getCustomEliminator_x3f___spec__9(x_2, x_37); +lean_dec(x_37); +lean_dec(x_2); +return x_38; +} +else +{ +size_t x_39; size_t x_40; uint64_t x_41; uint64_t x_42; size_t x_43; lean_object* x_44; lean_object* x_45; +x_39 = 0; +x_40 = lean_usize_of_nat(x_8); +lean_dec(x_8); +x_41 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_addCustomEliminatorEntry___spec__5(x_6, x_39, x_40, x_7); +lean_dec(x_6); +x_42 = lean_uint64_mix_hash(x_33, x_41); +x_43 = lean_hashmap_mk_idx(x_4, x_42); +x_44 = lean_array_uget(x_3, x_43); +lean_dec(x_3); +x_45 = l_Lean_AssocList_find_x3f___at_Lean_Meta_getCustomEliminator_x3f___spec__9(x_2, x_44); +lean_dec(x_44); +lean_dec(x_2); +return x_45; +} +} } } } @@ -12407,47 +13461,55 @@ return x_12; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_getCustomEliminator_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_EXPORT lean_object* l_Lean_Meta_getCustomEliminator_x3f___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, lean_object* x_8) { _start: { -lean_object* x_8; uint8_t x_9; -x_8 = lean_st_ref_get(x_6, x_7); -x_9 = !lean_is_exclusive(x_8); -if (x_9 == 0) +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_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_8, 0); -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -lean_dec(x_10); -x_12 = l_Lean_Meta_instInhabitedCustomEliminators; -x_13 = l_Lean_Meta_addCustomEliminator___closed__1; -x_14 = l_Lean_ScopedEnvExtension_getState___rarg(x_12, x_13, x_11); +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_11 = lean_ctor_get(x_9, 0); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); lean_dec(x_11); -x_15 = l_Lean_SMap_find_x3f___at_Lean_Meta_getCustomEliminator_x3f___spec__2(x_14, x_1); -lean_ctor_set(x_8, 0, x_15); -return x_8; +x_13 = l_Lean_Meta_instInhabitedCustomEliminators; +x_14 = l_Lean_Meta_addCustomEliminator___closed__1; +x_15 = l_Lean_ScopedEnvExtension_getState___rarg(x_13, x_14, x_12); +lean_dec(x_12); +x_16 = lean_box(x_1); +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 = l_Lean_SMap_find_x3f___at_Lean_Meta_getCustomEliminator_x3f___spec__2(x_15, x_17); +lean_ctor_set(x_9, 0, x_18); +return x_9; } 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; -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_ctor_get(x_16, 0); -lean_inc(x_18); -lean_dec(x_16); -x_19 = l_Lean_Meta_instInhabitedCustomEliminators; -x_20 = l_Lean_Meta_addCustomEliminator___closed__1; -x_21 = l_Lean_ScopedEnvExtension_getState___rarg(x_19, x_20, x_18); -lean_dec(x_18); -x_22 = l_Lean_SMap_find_x3f___at_Lean_Meta_getCustomEliminator_x3f___spec__2(x_21, x_1); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_17); -return x_23; +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_19 = lean_ctor_get(x_9, 0); +x_20 = lean_ctor_get(x_9, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_9); +x_21 = lean_ctor_get(x_19, 0); +lean_inc(x_21); +lean_dec(x_19); +x_22 = l_Lean_Meta_instInhabitedCustomEliminators; +x_23 = l_Lean_Meta_addCustomEliminator___closed__1; +x_24 = l_Lean_ScopedEnvExtension_getState___rarg(x_22, x_23, x_21); +lean_dec(x_21); +x_25 = lean_box(x_1); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_2); +x_27 = l_Lean_SMap_find_x3f___at_Lean_Meta_getCustomEliminator_x3f___spec__2(x_24, 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_20); +return x_28; } } } @@ -12463,106 +13525,106 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_getCustomEliminator_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_EXPORT lean_object* l_Lean_Meta_getCustomEliminator_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_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; lean_object* x_12; -x_7 = lean_box(0); -x_8 = lean_array_get_size(x_1); -x_9 = lean_usize_of_nat(x_8); -lean_dec(x_8); -x_10 = 0; -x_11 = l_Lean_Meta_getCustomEliminator_x3f___closed__1; +lean_object* x_8; lean_object* x_9; size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; +x_8 = lean_box(0); +x_9 = lean_array_get_size(x_1); +x_10 = lean_usize_of_nat(x_9); +lean_dec(x_9); +x_11 = 0; +x_12 = l_Lean_Meta_getCustomEliminator_x3f___closed__1; +lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -lean_inc(x_2); -x_12 = l_Array_forInUnsafe_loop___at_Lean_Meta_getCustomEliminator_x3f___spec__1(x_7, x_1, x_9, x_10, x_11, x_2, x_3, x_4, x_5, x_6); +x_13 = l_Array_forInUnsafe_loop___at_Lean_Meta_getCustomEliminator_x3f___spec__1(x_8, x_1, x_10, x_11, x_12, x_3, x_4, x_5, x_6, x_7); lean_dec(x_1); -if (lean_obj_tag(x_12) == 0) +if (lean_obj_tag(x_13) == 0) { -lean_object* x_13; lean_object* x_14; -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); +lean_object* x_14; lean_object* x_15; x_14 = lean_ctor_get(x_13, 0); lean_inc(x_14); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_15 = lean_ctor_get(x_12, 1); +x_15 = lean_ctor_get(x_14, 0); lean_inc(x_15); -lean_dec(x_12); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; x_16 = lean_ctor_get(x_13, 1); lean_inc(x_16); lean_dec(x_13); -x_17 = lean_box(0); -x_18 = l_Lean_Meta_getCustomEliminator_x3f___lambda__1(x_16, x_17, x_2, x_3, x_4, x_5, x_15); +x_17 = lean_ctor_get(x_14, 1); +lean_inc(x_17); +lean_dec(x_14); +x_18 = lean_box(0); +x_19 = l_Lean_Meta_getCustomEliminator_x3f___lambda__1(x_2, x_17, x_18, x_3, x_4, x_5, x_6, x_16); +lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -return x_18; +return x_19; } else { -uint8_t x_19; -lean_dec(x_13); +uint8_t x_20; +lean_dec(x_14); +lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -x_19 = !lean_is_exclusive(x_12); -if (x_19 == 0) +x_20 = !lean_is_exclusive(x_13); +if (x_20 == 0) { -lean_object* x_20; lean_object* x_21; -x_20 = lean_ctor_get(x_12, 0); -lean_dec(x_20); -x_21 = lean_ctor_get(x_14, 0); -lean_inc(x_21); -lean_dec(x_14); -lean_ctor_set(x_12, 0, x_21); -return x_12; +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_13, 0); +lean_dec(x_21); +x_22 = lean_ctor_get(x_15, 0); +lean_inc(x_22); +lean_dec(x_15); +lean_ctor_set(x_13, 0, x_22); +return x_13; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_12, 1); -lean_inc(x_22); -lean_dec(x_12); -x_23 = lean_ctor_get(x_14, 0); +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_13, 1); lean_inc(x_23); -lean_dec(x_14); -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_dec(x_13); +x_24 = lean_ctor_get(x_15, 0); +lean_inc(x_24); +lean_dec(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; } } } else { -uint8_t x_25; +uint8_t x_26; +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_12); -if (x_25 == 0) +x_26 = !lean_is_exclusive(x_13); +if (x_26 == 0) { -return x_12; +return x_13; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_12, 0); -x_27 = lean_ctor_get(x_12, 1); +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_13, 0); +x_28 = lean_ctor_get(x_13, 1); +lean_inc(x_28); lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_12); -x_28 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -return x_28; +lean_dec(x_13); +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; } } } @@ -12593,18 +13655,17 @@ x_8 = lean_box(x_7); return x_8; } } -LEAN_EXPORT lean_object* l_Array_isEqvAux___at_Lean_Meta_getCustomEliminator_x3f___spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Array_isEqvAux___at_Lean_Meta_getCustomEliminator_x3f___spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -uint8_t x_9; lean_object* x_10; -x_9 = l_Array_isEqvAux___at_Lean_Meta_getCustomEliminator_x3f___spec__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_3); +uint8_t x_7; lean_object* x_8; +x_7 = l_Array_isEqvAux___at_Lean_Meta_getCustomEliminator_x3f___spec__7(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_10 = lean_box(x_9); -return x_10; +x_8 = lean_box(x_7); +return x_8; } } LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_getCustomEliminator_x3f___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { @@ -12652,17 +13713,29 @@ lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_getCustomEliminator_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_EXPORT lean_object* l_Lean_Meta_getCustomEliminator_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_8; -x_8 = l_Lean_Meta_getCustomEliminator_x3f___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +uint8_t x_9; lean_object* x_10; +x_9 = lean_unbox(x_1); +lean_dec(x_1); +x_10 = l_Lean_Meta_getCustomEliminator_x3f___lambda__1(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); lean_dec(x_3); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_getCustomEliminator_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); -return x_8; +x_9 = l_Lean_Meta_getCustomEliminator_x3f(x_1, x_8, x_3, x_4, x_5, x_6, x_7); +return x_9; } } lean_object* initialize_Lean_Meta_Basic(uint8_t builtin, lean_object*); @@ -12892,32 +13965,24 @@ l_Lean_Meta_instInhabitedCustomEliminator___closed__1 = _init_l_Lean_Meta_instIn lean_mark_persistent(l_Lean_Meta_instInhabitedCustomEliminator___closed__1); l_Lean_Meta_instInhabitedCustomEliminator = _init_l_Lean_Meta_instInhabitedCustomEliminator(); lean_mark_persistent(l_Lean_Meta_instInhabitedCustomEliminator); -l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__1 = _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__1); -l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__2 = _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__2(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__2); -l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__3 = _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__3(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__3); -l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__4 = _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__4(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__4); -l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__5 = _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__5(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__5); -l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__6 = _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__6(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__6); -l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__7 = _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__7(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__7); -l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__8 = _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__8(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__8); -l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__9 = _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__9(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__9); -l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__10 = _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__10(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__10); -l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__11 = _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__11(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__11); -l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__12 = _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__12(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__12); -l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__13 = _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__13(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2148____closed__13); +l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__1 = _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__1); +l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__2 = _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__2); +l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__3 = _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__3); +l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__4 = _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__4); +l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__5 = _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__5); +l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__6 = _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__6); +l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__7 = _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__7); +l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__8 = _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__8(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__8); +l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__9 = _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__9(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminator____x40_Lean_Meta_Tactic_ElimInfo___hyg_2157____closed__9); l_Lean_Meta_instReprCustomEliminator___closed__1 = _init_l_Lean_Meta_instReprCustomEliminator___closed__1(); lean_mark_persistent(l_Lean_Meta_instReprCustomEliminator___closed__1); l_Lean_Meta_instReprCustomEliminator = _init_l_Lean_Meta_instReprCustomEliminator(); @@ -12932,6 +13997,16 @@ l_Lean_Meta_CustomEliminators_map___default___closed__4 = _init_l_Lean_Meta_Cust lean_mark_persistent(l_Lean_Meta_CustomEliminators_map___default___closed__4); l_Lean_Meta_CustomEliminators_map___default___closed__5 = _init_l_Lean_Meta_CustomEliminators_map___default___closed__5(); lean_mark_persistent(l_Lean_Meta_CustomEliminators_map___default___closed__5); +l_Lean_Meta_CustomEliminators_map___default___closed__6 = _init_l_Lean_Meta_CustomEliminators_map___default___closed__6(); +lean_mark_persistent(l_Lean_Meta_CustomEliminators_map___default___closed__6); +l_Lean_Meta_CustomEliminators_map___default___closed__7 = _init_l_Lean_Meta_CustomEliminators_map___default___closed__7(); +lean_mark_persistent(l_Lean_Meta_CustomEliminators_map___default___closed__7); +l_Lean_Meta_CustomEliminators_map___default___closed__8 = _init_l_Lean_Meta_CustomEliminators_map___default___closed__8(); +lean_mark_persistent(l_Lean_Meta_CustomEliminators_map___default___closed__8); +l_Lean_Meta_CustomEliminators_map___default___closed__9 = _init_l_Lean_Meta_CustomEliminators_map___default___closed__9(); +lean_mark_persistent(l_Lean_Meta_CustomEliminators_map___default___closed__9); +l_Lean_Meta_CustomEliminators_map___default___closed__10 = _init_l_Lean_Meta_CustomEliminators_map___default___closed__10(); +lean_mark_persistent(l_Lean_Meta_CustomEliminators_map___default___closed__10); l_Lean_Meta_CustomEliminators_map___default = _init_l_Lean_Meta_CustomEliminators_map___default(); lean_mark_persistent(l_Lean_Meta_CustomEliminators_map___default); l_Lean_Meta_instInhabitedCustomEliminators___closed__1 = _init_l_Lean_Meta_instInhabitedCustomEliminators___closed__1(); @@ -12940,72 +14015,73 @@ l_Lean_Meta_instInhabitedCustomEliminators___closed__2 = _init_l_Lean_Meta_instI lean_mark_persistent(l_Lean_Meta_instInhabitedCustomEliminators___closed__2); l_Lean_Meta_instInhabitedCustomEliminators = _init_l_Lean_Meta_instInhabitedCustomEliminators(); lean_mark_persistent(l_Lean_Meta_instInhabitedCustomEliminators); -l_Lean_SMap_toList___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__1___closed__1 = _init_l_Lean_SMap_toList___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__1___closed__1(); -lean_mark_persistent(l_Lean_SMap_toList___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__1___closed__1); -l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__1 = _init_l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__1(); -lean_mark_persistent(l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__1); -l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__2 = _init_l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__2(); -lean_mark_persistent(l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__2); -l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__3 = _init_l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__3(); -lean_mark_persistent(l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__3); -l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__4 = _init_l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__4(); -lean_mark_persistent(l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__4); -l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__5 = _init_l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__5(); -lean_mark_persistent(l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__5); -l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__6 = _init_l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__6(); -lean_mark_persistent(l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__6); -l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__7 = _init_l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__7(); -lean_mark_persistent(l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__10___closed__7); -l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8___closed__1 = _init_l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8___closed__1(); -lean_mark_persistent(l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8___closed__1); -l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8___closed__2 = _init_l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8___closed__2(); -lean_mark_persistent(l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8___closed__2); -l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8___closed__3 = _init_l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8___closed__3(); -lean_mark_persistent(l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8___closed__3); -l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8___closed__4 = _init_l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8___closed__4(); -lean_mark_persistent(l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8___closed__4); -l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8___closed__5 = _init_l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8___closed__5(); -lean_mark_persistent(l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8___closed__5); -l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8___closed__6 = _init_l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8___closed__6(); -lean_mark_persistent(l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____spec__8___closed__6); -l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__1 = _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__1); -l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__2 = _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__2(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__2); -l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__3 = _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__3(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__3); -l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__4 = _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__4(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__4); -l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__5 = _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__5(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__5); -l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__6 = _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__6(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__6); -l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__7 = _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__7(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2237____closed__7); +l_Lean_SMap_toList___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__1___closed__1 = _init_l_Lean_SMap_toList___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__1___closed__1(); +lean_mark_persistent(l_Lean_SMap_toList___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__1___closed__1); +l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___closed__1 = _init_l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___closed__1(); +lean_mark_persistent(l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___closed__1); +l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___closed__2 = _init_l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___closed__2(); +lean_mark_persistent(l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___closed__2); +l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___closed__3 = _init_l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___closed__3(); +lean_mark_persistent(l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___closed__3); +l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___closed__4 = _init_l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___closed__4(); +lean_mark_persistent(l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___closed__4); +l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___closed__5 = _init_l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___closed__5(); +lean_mark_persistent(l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___closed__5); +l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___closed__6 = _init_l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___closed__6(); +lean_mark_persistent(l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__11___closed__6); +l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8___closed__1 = _init_l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8___closed__1(); +lean_mark_persistent(l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8___closed__1); +l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8___closed__2 = _init_l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8___closed__2(); +lean_mark_persistent(l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8___closed__2); +l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8___closed__3 = _init_l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8___closed__3(); +lean_mark_persistent(l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8___closed__3); +l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8___closed__4 = _init_l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8___closed__4(); +lean_mark_persistent(l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8___closed__4); +l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8___closed__5 = _init_l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8___closed__5(); +lean_mark_persistent(l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8___closed__5); +l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8___closed__6 = _init_l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8___closed__6(); +lean_mark_persistent(l_List_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____spec__8___closed__6); +l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____closed__1 = _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____closed__1); +l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____closed__2 = _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____closed__2); +l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____closed__3 = _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____closed__3); +l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____closed__4 = _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____closed__4); +l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____closed__5 = _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____closed__5); +l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____closed__6 = _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____closed__6); +l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____closed__7 = _init_l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2270____closed__7); l_Lean_Meta_instReprCustomEliminators___closed__1 = _init_l_Lean_Meta_instReprCustomEliminators___closed__1(); lean_mark_persistent(l_Lean_Meta_instReprCustomEliminators___closed__1); l_Lean_Meta_instReprCustomEliminators = _init_l_Lean_Meta_instReprCustomEliminators(); lean_mark_persistent(l_Lean_Meta_instReprCustomEliminators); l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__1 = _init_l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__1(); +l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__2 = _init_l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__2(); +l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__3 = _init_l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__3(); +l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__4 = _init_l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4___closed__4(); l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3___closed__1 = _init_l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3___closed__1(); l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3___closed__2 = _init_l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3___closed__2(); l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3___closed__3 = _init_l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3___closed__3(); lean_mark_persistent(l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addCustomEliminatorEntry___spec__3___closed__3); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__1 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__1(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__1); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__2 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__2(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__2); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__3 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__3(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__3); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__4 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__4(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__4); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__5 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__5(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__5); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__6 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__6(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__6); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__7 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__7(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298____closed__7); -if (builtin) {res = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2298_(lean_io_mk_world()); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____closed__1 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____closed__1(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____closed__1); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____closed__2 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____closed__2(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____closed__2); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____closed__3 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____closed__3(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____closed__3); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____closed__4 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____closed__4(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____closed__4); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____closed__5 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____closed__5(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____closed__5); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____closed__6 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____closed__6(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____closed__6); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____closed__7 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____closed__7(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335____closed__7); +if (builtin) {res = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2335_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Meta_customEliminatorExt = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Meta_customEliminatorExt); @@ -13062,77 +14138,94 @@ l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addCustomEliminator___spec__1___clo lean_mark_persistent(l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addCustomEliminator___spec__1___closed__17); l_Lean_Meta_addCustomEliminator___closed__1 = _init_l_Lean_Meta_addCustomEliminator___closed__1(); lean_mark_persistent(l_Lean_Meta_addCustomEliminator___closed__1); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__1 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__1(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__1); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__2 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__2(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__2); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__3 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__3(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__3); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__4 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__4(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__4); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__5 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__5(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__5); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__6 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__6(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__6); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__7 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__7(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__7); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__8 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__8(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__8); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__9 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__9(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__9); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__10 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__10(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__10); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__11 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__11(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__1___closed__11); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__2___closed__1 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__2___closed__1(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__2___closed__1); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__2___closed__2 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__2___closed__2(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____lambda__2___closed__2); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__1 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__1(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__1); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__2 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__2(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__2); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__3 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__3(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__3); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__4 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__4(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__4); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__5 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__5(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__5); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__6 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__6(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__6); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__7 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__7(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__7); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__8 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__8(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__8); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__9 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__9(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__9); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__10 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__10(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__10); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__11 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__11(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__11); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__12 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__12(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__12); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__13 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__13(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__13); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__14 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__14(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__14); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__15 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__15(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__15); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__16 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__16(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__16); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__17 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__17(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__17); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__18 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__18(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__18); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__19 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__19(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__19); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__20 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__20(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__20); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__21 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__21(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__21); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__22 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__22(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946____closed__22); -if (builtin) {res = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2946_(lean_io_mk_world()); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__1 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__1); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__2 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__2); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__3 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__3(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__3); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__4 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__4(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__4); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__5 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__5(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__5); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__6 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__6(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__6); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__7 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__7(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__7); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__8 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__8(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__8); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__9 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__9(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__9); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__10 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__10(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__10); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__11 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__11(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__11); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__2___closed__1 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__2___closed__1); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__2___closed__2 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__2___closed__2); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__1 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__1(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__1); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__2 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__2(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__2); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__3 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__3(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__3); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__4 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__4(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__4); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__5 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__5(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__5); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__6 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__6(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__6); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__7 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__7(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__7); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__8 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__8(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__8); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__9 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__9(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__9); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__10 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__10(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__10); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__11 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__11(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__11); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__12 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__12(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__12); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__13 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__13(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__13); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__14 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__14(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__14); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__15 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__15(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__15); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__16 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__16(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__16); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__17 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__17(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__17); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__18 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__18(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__18); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__19 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__19(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__19); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__20 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__20(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__20); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__21 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__21(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__21); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__22 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__22(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____closed__22); +if (builtin) {res = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____closed__1 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____closed__1(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____closed__1); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____closed__2 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____closed__2(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____closed__2); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____closed__3 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____closed__3(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____closed__3); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____closed__4 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____closed__4(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____closed__4); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____closed__5 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____closed__5(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____closed__5); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____closed__6 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____closed__6(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____closed__6); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____closed__7 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____closed__7(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____closed__7); +if (builtin) {res = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }l_Array_forInUnsafe_loop___at_Lean_Meta_getCustomEliminator_x3f___spec__1___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Meta_getCustomEliminator_x3f___spec__1___closed__1(); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Simp/Main.c b/stage0/stdlib/Lean/Meta/Tactic/Simp/Main.c index 4e2ba55d6adb..83491d47da28 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Simp/Main.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Simp/Main.c @@ -61,6 +61,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_dsimp(lean_object*, lean_object*, lean_obje static lean_object* l_Lean_Meta_Simp_simpArrow___lambda__4___closed__8; lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Main___hyg_5____closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Simp_isOfScientificLit___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simpLoop_visitPostContinue___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceStep___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_Meta_Simp_simpArrow___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*); @@ -99,6 +100,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit___at___private_Lean_Meta_Ta static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_trySimpCongrTheorem_x3f___spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_applySimpResultToLocalDecl___boxed(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*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_doNotVisitOfScientific___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_Simp_simpForall___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_Array_forInUnsafe_loop___at_Lean_Meta_dsimpGoal___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*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_doNotVisitOfNat___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -121,6 +123,7 @@ 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_Lean_Meta_simpGoal(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_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceFVar(lean_object*, lean_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_doNotVisit___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_withLetDecl___at___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___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*); lean_object* l_Lean_mkHashSetImp___rarg(lean_object*); lean_object* l_Lean_Core_checkSystem(lean_object*, lean_object*, lean_object*, lean_object*); @@ -218,6 +221,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Simp_visitFn___lambda__1___boxed(lean_objec LEAN_EXPORT lean_object* l_Lean_isProjectionFn___at___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_unfold_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_Meta_Simp_simpProj___spec__11___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_simpGoal___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_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___lambda__1___closed__3; lean_object* l_Lean_Expr_appArg_x21(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_dsimpGoal___lambda__6(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_ConstantInfo_value_x3f(lean_object*); @@ -351,6 +355,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___lambda__5___ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_simpGoal___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_EXPORT lean_object* l_List_forIn_loop___at_Lean_Meta_Simp_congr___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_Array_forInUnsafe_loop___at_Lean_Meta_simpGoal___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_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___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* l_Lean_Meta_mkAppM(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simpLambda___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_Meta_dsimpGoal___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*); @@ -382,6 +387,7 @@ extern lean_object* l_Lean_instInhabitedExpr; LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit_visitLet___at___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___spec__6___lambda__1___boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_getSimpLetCase___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_throwCongrHypothesisFailed___rarg___closed__2; +static lean_object* l_Lean_Meta_Simp_isOfScientificLit___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simpForall___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_Meta_Simp_simpArrow___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_Simp_trySimpCongrTheorem_x3f___lambda__3___closed__1; @@ -486,6 +492,7 @@ lean_object* l_Lean_Meta_mkOfEqTrue(lean_object*, lean_object*, lean_object*, le uint8_t l_Lean_Expr_hasAnyFVar_visit___at_Lean_Expr_containsFVar___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Expr_bindingDomain_x21(lean_object*); lean_object* l_Lean_Expr_letValue_x21(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_doNotVisitOfScientific(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_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___spec__12___boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_Meta_simpGoal___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_trySimpCongrTheorem_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*, lean_object*, lean_object*, lean_object*); @@ -509,6 +516,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_simpTargetCore___lambda__1___boxed(lean_obj lean_object* l_Lean_Meta_Simp_recordSimpTheorem(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_transform___at___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___spec__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___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_doNotVisit(lean_object*, lean_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_simpArrow___lambda__3___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___spec__8___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_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_Simp_simpProj___spec__7(lean_object*, lean_object*, size_t, size_t); @@ -570,7 +578,7 @@ lean_object* lean_nat_mul(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___spec__13___boxed(lean_object**); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_Meta_Simp_simpProj___spec__3(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkFunExt(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_dsimpImpl___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___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_Simp_simpLoop___closed__1; static lean_object* l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___lambda__3___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_simpStep(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*); @@ -600,6 +608,7 @@ static lean_object* l_Lean_Meta_Simp_SimpLetCase_noConfusion___rarg___closed__1; lean_object* l_Lean_Meta_mkCongrFun(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isTrue(lean_object*); lean_object* l_Lean_MVarId_tryClearMany(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_Meta_Simp_isOfScientificLit(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___lambda__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*); 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*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_unfold_x3f_unfoldDeclToUnfold_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -619,6 +628,7 @@ lean_object* l_Lean_Expr_betaRev(lean_object*, lean_object*, uint8_t, uint8_t); uint8_t l_Lean_Meta_SimpTheoremsArray_isLetDeclToUnfold(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_Simp_isOfScientificLit___closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_Meta_Simp_simpProj___spec__13___boxed(lean_object*, lean_object*); lean_object* l_Lean_instantiateMVars___at_Lean_Meta_Simp_synthesizeArgs___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_applySimpResultToProp(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -670,6 +680,7 @@ lean_object* l_Lean_Meta_Simp_getConfig___rarg(lean_object*, lean_object*, lean_ static lean_object* l_Lean_Meta_applySimpResultToProp___closed__1; 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___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_doNotVisitOfScientific___closed__1; static lean_object* l_Lean_Meta_Simp_simpArrow___lambda__4___closed__14; lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Simp_congrArgs___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_simpGoal___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*); @@ -704,11 +715,13 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_simpGoal___spe lean_object* l_Lean_MessageData_ofName(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceProjFn_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_EXPORT lean_object* l_Lean_Meta_withLetDecl___at___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___spec__9(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Simp_isOfScientificLit___closed__1; lean_object* lean_expr_instantiate1(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_simpArrow___lambda__4___closed__12; static lean_object* l_Lean_Meta_Simp_simpProj___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___spec__9___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_simpImpl_go___lambda__2___closed__2; +static lean_object* l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___lambda__2___closed__1; lean_object* l_Lean_Meta_SimpTheoremsArray_eraseTheorem(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_throwCongrHypothesisFailed___rarg___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_dsimpGoal___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*); @@ -729,6 +742,7 @@ static lean_object* l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___closed__1; lean_object* l_Lean_Expr_letBody_x21(lean_object*); lean_object* l_Lean_Expr_getNumHeadLambdas(lean_object*); static lean_object* l_Lean_Meta_Simp_simpForall___lambda__4___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_doNotVisitOfNat___closed__1; static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Main___hyg_5____closed__1() { _start: { @@ -1148,6 +1162,91 @@ lean_dec(x_2); return x_10; } } +static lean_object* _init_l_Lean_Meta_Simp_isOfScientificLit___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("OfScientific", 12); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Simp_isOfScientificLit___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("ofScientific", 12); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Simp_isOfScientificLit___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Simp_isOfScientificLit___closed__1; +x_2 = l_Lean_Meta_Simp_isOfScientificLit___closed__2; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT uint8_t l_Lean_Meta_Simp_isOfScientificLit(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; uint8_t x_4; +x_2 = l_Lean_Meta_Simp_isOfScientificLit___closed__3; +x_3 = lean_unsigned_to_nat(5u); +x_4 = l_Lean_Expr_isAppOfArity(x_1, x_2, x_3); +if (x_4 == 0) +{ +uint8_t x_5; +x_5 = 0; +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_6 = lean_unsigned_to_nat(0u); +x_7 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_6); +x_8 = lean_unsigned_to_nat(4u); +x_9 = lean_nat_sub(x_7, x_8); +x_10 = lean_unsigned_to_nat(1u); +x_11 = lean_nat_sub(x_9, x_10); +lean_dec(x_9); +x_12 = l_Lean_Expr_getRevArg_x21(x_1, x_11); +x_13 = l_Lean_Expr_isRawNatLit(x_12); +lean_dec(x_12); +if (x_13 == 0) +{ +uint8_t x_14; +lean_dec(x_7); +x_14 = 0; +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_15 = lean_unsigned_to_nat(2u); +x_16 = lean_nat_sub(x_7, x_15); +lean_dec(x_7); +x_17 = lean_nat_sub(x_16, x_10); +lean_dec(x_16); +x_18 = l_Lean_Expr_getRevArg_x21(x_1, x_17); +x_19 = l_Lean_Expr_isRawNatLit(x_18); +lean_dec(x_18); +return x_19; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Simp_isOfScientificLit___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Lean_Meta_Simp_isOfScientificLit(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} static lean_object* _init_l_Lean_getProjectionFnInfo_x3f___at___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceProjFn_x3f___spec__1___closed__1() { _start: { @@ -19109,7 +19208,7 @@ static lean_object* _init_l_Lean_Meta_Simp_simpLet___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_Simp_simpLet___closed__1; x_2 = l_Lean_Meta_Simp_simpLet___closed__2; -x_3 = lean_unsigned_to_nat(376u); +x_3 = lean_unsigned_to_nat(380u); x_4 = lean_unsigned_to_nat(29u); x_5 = l_Lean_Meta_Simp_simpLet___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -19679,53 +19778,87 @@ lean_dec(x_3); return x_12; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_doNotVisitOfNat(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Tactic_Simp_Main_0__Lean_Meta_Simp_doNotVisit(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_10; -lean_inc(x_1); -x_10 = l_Lean_Meta_Simp_isOfNatNatLit(x_1); -if (x_10 == 0) +lean_object* x_12; uint8_t x_13; +lean_inc(x_3); +x_12 = lean_apply_1(x_1, x_3); +x_13 = lean_unbox(x_12); +lean_dec(x_12); +if (x_13 == 0) { -lean_object* x_11; lean_object* x_12; lean_object* x_13; -lean_dec(x_3); -x_11 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_11, 0, x_1); -x_12 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_12, 0, x_11); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_12); -lean_ctor_set(x_13, 1, x_9); -return x_13; +lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_dec(x_5); +lean_dec(x_2); +x_14 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_14, 0, x_3); +x_15 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_15, 0, 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_11); +return x_16; } else { -lean_object* x_14; uint8_t x_15; -x_14 = l_Lean_Meta_Simp_isOfNatNatLit___closed__3; -x_15 = l_Lean_Meta_Simp_Context_isDeclToUnfold(x_3, x_14); -if (x_15 == 0) +uint8_t x_17; +x_17 = l_Lean_Meta_Simp_Context_isDeclToUnfold(x_5, x_2); +if (x_17 == 0) { -lean_object* x_16; lean_object* x_17; -x_16 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_16, 0, x_1); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set(x_17, 1, x_9); -return x_17; +lean_object* x_18; lean_object* x_19; +x_18 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_18, 0, x_3); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_11); +return x_19; } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_18, 0, x_1); -x_19 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_19, 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_9); -return x_20; +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_3); +x_21 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_21, 0, 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_11); +return x_22; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_doNotVisit___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_doNotVisit(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_4); +return x_12; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_doNotVisitOfNat___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Simp_isOfNatNatLit___boxed), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_doNotVisitOfNat(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_doNotVisitOfNat___closed__1; +x_11 = l_Lean_Meta_Simp_isOfNatNatLit___closed__3; +x_12 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_doNotVisit(x_10, x_11, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_12; } } LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_doNotVisitOfNat___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { @@ -19742,6 +19875,38 @@ lean_dec(x_2); return x_10; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_doNotVisitOfScientific___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Simp_isOfScientificLit___boxed), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_doNotVisitOfScientific(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_doNotVisitOfScientific___closed__1; +x_11 = l_Lean_Meta_Simp_isOfScientificLit___closed__3; +x_12 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_doNotVisit(x_10, x_11, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_12; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_doNotVisitOfScientific___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_doNotVisitOfScientific(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_2); +return x_10; +} +} LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit_visitPost___at___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___spec__3(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, 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: { @@ -22594,32 +22759,59 @@ static lean_object* _init_l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Si _start: { lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_doNotVisitOfScientific___boxed), 9, 0); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___lambda__1___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___lambda__1___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___lambda__1___closed__2; +x_3 = lean_alloc_closure((void*)(l_Lean_Meta_Simp_dandThen), 11, 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___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___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; +x_11 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___lambda__1___closed__3; +x_12 = l_Lean_Meta_Simp_dandThen(x_1, x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_12; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpReduce), 9, 0); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; uint8_t x_19; lean_object* x_20; +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; uint8_t x_18; lean_object* x_19; lean_dec(x_3); x_12 = lean_ctor_get(x_4, 2); lean_inc(x_12); -x_13 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___lambda__1___closed__1; -x_14 = lean_alloc_closure((void*)(l_Lean_Meta_Simp_dandThen), 11, 2); -lean_closure_set(x_14, 0, x_12); -lean_closure_set(x_14, 1, x_13); -x_15 = lean_ctor_get(x_4, 3); -lean_inc(x_15); -x_16 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___lambda__1___closed__2; -x_17 = lean_alloc_closure((void*)(l_Lean_Meta_Simp_dandThen), 11, 2); -lean_closure_set(x_17, 0, x_15); -lean_closure_set(x_17, 1, x_16); -x_18 = lean_ctor_get_uint8(x_1, sizeof(void*)*2 + 3); +x_13 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___lambda__1), 10, 1); +lean_closure_set(x_13, 0, x_12); +x_14 = lean_ctor_get(x_4, 3); +lean_inc(x_14); +x_15 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___lambda__2___closed__1; +x_16 = lean_alloc_closure((void*)(l_Lean_Meta_Simp_dandThen), 11, 2); +lean_closure_set(x_16, 0, x_14); +lean_closure_set(x_16, 1, x_15); +x_17 = lean_ctor_get_uint8(x_1, sizeof(void*)*2 + 3); lean_dec(x_1); -x_19 = 0; -x_20 = l_Lean_Meta_transform___at___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___spec__1(x_2, x_14, x_17, x_18, x_19, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_20; +x_18 = 0; +x_19 = l_Lean_Meta_transform___at___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___spec__1(x_2, x_13, x_16, x_17, x_18, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_19; } } LEAN_EXPORT lean_object* lean_dsimp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { @@ -22669,7 +22861,7 @@ x_17 = lean_ctor_get(x_10, 1); lean_inc(x_17); lean_dec(x_10); x_18 = lean_box(0); -x_19 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___lambda__1(x_11, x_1, x_18, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_17); +x_19 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___lambda__2(x_11, x_1, x_18, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_17); return x_19; } } @@ -27543,13 +27735,17 @@ lean_inc(x_1); x_10 = l_Lean_Meta_Simp_isOfNatNatLit(x_1); if (x_10 == 0) { -lean_object* x_11; -x_11 = l_Lean_Meta_Simp_congr(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_11; +uint8_t x_11; +x_11 = l_Lean_Meta_Simp_isOfScientificLit(x_1); +if (x_11 == 0) +{ +lean_object* x_12; +x_12 = l_Lean_Meta_Simp_congr(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_12; } else { -lean_object* x_12; uint32_t x_13; uint8_t x_14; lean_object* x_15; lean_object* x_16; +lean_object* x_13; uint32_t x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -27557,18 +27753,42 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_12 = lean_box(0); -x_13 = 0; -x_14 = 1; -x_15 = lean_alloc_ctor(0, 2, 5); -lean_ctor_set(x_15, 0, x_1); -lean_ctor_set(x_15, 1, x_12); -lean_ctor_set_uint32(x_15, sizeof(void*)*2, x_13); -lean_ctor_set_uint8(x_15, sizeof(void*)*2 + 4, x_14); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_9); -return x_16; +x_13 = lean_box(0); +x_14 = 0; +x_15 = 1; +x_16 = lean_alloc_ctor(0, 2, 5); +lean_ctor_set(x_16, 0, x_1); +lean_ctor_set(x_16, 1, x_13); +lean_ctor_set_uint32(x_16, sizeof(void*)*2, x_14); +lean_ctor_set_uint8(x_16, sizeof(void*)*2 + 4, x_15); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_9); +return x_17; +} +} +else +{ +lean_object* x_18; uint32_t x_19; uint8_t x_20; lean_object* x_21; lean_object* 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_2); +x_18 = lean_box(0); +x_19 = 0; +x_20 = 1; +x_21 = lean_alloc_ctor(0, 2, 5); +lean_ctor_set(x_21, 0, x_1); +lean_ctor_set(x_21, 1, x_18); +lean_ctor_set_uint32(x_21, sizeof(void*)*2, x_19); +lean_ctor_set_uint8(x_21, sizeof(void*)*2 + 4, 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_9); +return x_22; } } } @@ -27586,7 +27806,7 @@ static lean_object* _init_l_Lean_Meta_Simp_simpStep___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_Meta_Simp_simpLet___closed__1; x_2 = l_Lean_Meta_Simp_simpStep___closed__1; -x_3 = lean_unsigned_to_nat(568u); +x_3 = lean_unsigned_to_nat(580u); x_4 = lean_unsigned_to_nat(20u); x_5 = l_Lean_Meta_Simp_simpLet___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -43443,6 +43663,12 @@ l_Lean_Meta_Simp_isOfNatNatLit___closed__2 = _init_l_Lean_Meta_Simp_isOfNatNatLi lean_mark_persistent(l_Lean_Meta_Simp_isOfNatNatLit___closed__2); l_Lean_Meta_Simp_isOfNatNatLit___closed__3 = _init_l_Lean_Meta_Simp_isOfNatNatLit___closed__3(); lean_mark_persistent(l_Lean_Meta_Simp_isOfNatNatLit___closed__3); +l_Lean_Meta_Simp_isOfScientificLit___closed__1 = _init_l_Lean_Meta_Simp_isOfScientificLit___closed__1(); +lean_mark_persistent(l_Lean_Meta_Simp_isOfScientificLit___closed__1); +l_Lean_Meta_Simp_isOfScientificLit___closed__2 = _init_l_Lean_Meta_Simp_isOfScientificLit___closed__2(); +lean_mark_persistent(l_Lean_Meta_Simp_isOfScientificLit___closed__2); +l_Lean_Meta_Simp_isOfScientificLit___closed__3 = _init_l_Lean_Meta_Simp_isOfScientificLit___closed__3(); +lean_mark_persistent(l_Lean_Meta_Simp_isOfScientificLit___closed__3); l_Lean_getProjectionFnInfo_x3f___at___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceProjFn_x3f___spec__1___closed__1 = _init_l_Lean_getProjectionFnInfo_x3f___at___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceProjFn_x3f___spec__1___closed__1(); lean_mark_persistent(l_Lean_getProjectionFnInfo_x3f___at___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceProjFn_x3f___spec__1___closed__1); l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceProjFn_x3f___lambda__1___closed__1 = _init_l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceProjFn_x3f___lambda__1___closed__1(); @@ -43597,12 +43823,20 @@ l_Lean_Meta_Simp_simpLet___closed__3 = _init_l_Lean_Meta_Simp_simpLet___closed__ lean_mark_persistent(l_Lean_Meta_Simp_simpLet___closed__3); l_Lean_Meta_Simp_simpLet___closed__4 = _init_l_Lean_Meta_Simp_simpLet___closed__4(); lean_mark_persistent(l_Lean_Meta_Simp_simpLet___closed__4); +l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_doNotVisitOfNat___closed__1 = _init_l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_doNotVisitOfNat___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_doNotVisitOfNat___closed__1); +l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_doNotVisitOfScientific___closed__1 = _init_l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_doNotVisitOfScientific___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_doNotVisitOfScientific___closed__1); l_Lean_Meta_transform___at___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___spec__1___closed__1 = _init_l_Lean_Meta_transform___at___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___spec__1___closed__1(); lean_mark_persistent(l_Lean_Meta_transform___at___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___spec__1___closed__1); l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___lambda__1___closed__1 = _init_l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___lambda__1___closed__1(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___lambda__1___closed__1); l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___lambda__1___closed__2 = _init_l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___lambda__1___closed__2(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___lambda__1___closed__2); +l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___lambda__1___closed__3 = _init_l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___lambda__1___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___lambda__1___closed__3); +l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___lambda__2___closed__1 = _init_l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___lambda__2___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___lambda__2___closed__1); l_Lean_Expr_withAppAux___at_Lean_Meta_Simp_processCongrHypothesis___spec__1___lambda__3___closed__1 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Simp_processCongrHypothesis___spec__1___lambda__3___closed__1(); lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_Simp_processCongrHypothesis___spec__1___lambda__3___closed__1); l_Lean_Expr_withAppAux___at_Lean_Meta_Simp_processCongrHypothesis___spec__1___lambda__3___closed__2 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Simp_processCongrHypothesis___spec__1___lambda__3___closed__2(); diff --git a/tests/lean/inductionErrors.lean.expected.out b/tests/lean/inductionErrors.lean.expected.out index 502af53f1550..5e7fb7096817 100644 --- a/tests/lean/inductionErrors.lean.expected.out +++ b/tests/lean/inductionErrors.lean.expected.out @@ -18,11 +18,11 @@ x y : Nat ⊢ 0 + (y + 1) = y + 1 inductionErrors.lean:40:14-40:18: error: unsolved goals case zero -⊢ 0 + Nat.zero = Nat.zero +⊢ 0 + 0 = 0 inductionErrors.lean:41:14-41:18: error: unsolved goals case succ y : Nat -⊢ 0 + Nat.succ y = Nat.succ y +⊢ 0 + (y + 1) = y + 1 inductionErrors.lean:50:2-50:16: error: alternative 'cons' is not needed inductionErrors.lean:55:2-55:16: error: alternative 'cons' is not needed inductionErrors.lean:60:2-60:40: error: invalid alternative name 'upper2' diff --git a/tests/lean/interactive/anonHyp.lean.expected.out b/tests/lean/interactive/anonHyp.lean.expected.out index a60251092bdb..143623916904 100644 --- a/tests/lean/interactive/anonHyp.lean.expected.out +++ b/tests/lean/interactive/anonHyp.lean.expected.out @@ -1,5 +1,5 @@ {"textDocument": {"uri": "file:///anonHyp.lean"}, "position": {"line": 2, "character": 4}} {"rendered": - "```lean\ncase succ\nn✝ : Nat\nn_ih✝ : n✝ + 1 > 0\n⊢ Nat.succ n✝ + 1 > 0\n```", - "goals": ["case succ\nn✝ : Nat\nn_ih✝ : n✝ + 1 > 0\n⊢ Nat.succ n✝ + 1 > 0"]} + "```lean\ncase succ\nn✝ : Nat\na✝ : n✝ + 1 > 0\n⊢ n✝ + 1 + 1 > 0\n```", + "goals": ["case succ\nn✝ : Nat\na✝ : n✝ + 1 > 0\n⊢ n✝ + 1 + 1 > 0"]} diff --git a/tests/lean/interactive/goalIssue.lean.expected.out b/tests/lean/interactive/goalIssue.lean.expected.out index 83627f6f41b7..a0eb9a2cdd1d 100644 --- a/tests/lean/interactive/goalIssue.lean.expected.out +++ b/tests/lean/interactive/goalIssue.lean.expected.out @@ -4,5 +4,5 @@ "goals": ["x : Nat\nthis : x + x = x + x\n⊢ 0 + x = x"]} {"textDocument": {"uri": "file:///goalIssue.lean"}, "position": {"line": 8, "character": 12}} -{"rendered": "```lean\ncase zero\n⊢ 0 + Nat.zero = Nat.zero\n```", - "goals": ["case zero\n⊢ 0 + Nat.zero = Nat.zero"]} +{"rendered": "```lean\ncase zero\n⊢ 0 + 0 = 0\n```", + "goals": ["case zero\n⊢ 0 + 0 = 0"]} diff --git a/tests/lean/interactive/hover.lean.expected.out b/tests/lean/interactive/hover.lean.expected.out index 5437b2017f25..6f0b0b394a55 100644 --- a/tests/lean/interactive/hover.lean.expected.out +++ b/tests/lean/interactive/hover.lean.expected.out @@ -452,18 +452,19 @@ {"textDocument": {"uri": "file:///hover.lean"}, "position": {"line": 235, "character": 4}} {"range": - {"start": {"line": 235, "character": 4}, "end": {"line": 235, "character": 8}}, + {"start": {"line": 235, "character": 2}, "end": {"line": 235, "character": 8}}, "contents": {"value": - "```lean\nNat.zero : ℕ\n```\n***\n`Nat.zero`, normally written `0 : Nat`, is the smallest natural number.\nThis is one of the two constructors of `Nat`. \n***\n*import Init.Prelude*", + "The left hand side of an induction arm, `| foo a b c` or `| @foo a b c`\nwhere `foo` is a constructor of the inductive type and `a b c` are the arguments\nto the constructor.\n", "kind": "markdown"}} {"textDocument": {"uri": "file:///hover.lean"}, "position": {"line": 238, "character": 4}} {"range": - {"start": {"line": 238, "character": 4}, "end": {"line": 238, "character": 8}}, + {"start": {"line": 238, "character": 2}, + "end": {"line": 238, "character": 10}}, "contents": {"value": - "```lean\nNat.succ (n : ℕ) : ℕ\n```\n***\nThe successor function on natural numbers, `succ n = n + 1`.\nThis is one of the two constructors of `Nat`. \n***\n*import Init.Prelude*", + "The left hand side of an induction arm, `| foo a b c` or `| @foo a b c`\nwhere `foo` is a constructor of the inductive type and `a b c` are the arguments\nto the constructor.\n", "kind": "markdown"}} {"textDocument": {"uri": "file:///hover.lean"}, "position": {"line": 238, "character": 9}} @@ -485,10 +486,10 @@ {"textDocument": {"uri": "file:///hover.lean"}, "position": {"line": 246, "character": 4}} {"range": - {"start": {"line": 246, "character": 4}, "end": {"line": 246, "character": 8}}, + {"start": {"line": 246, "character": 2}, "end": {"line": 246, "character": 8}}, "contents": {"value": - "```lean\nNat.zero : ℕ\n```\n***\n`Nat.zero`, normally written `0 : Nat`, is the smallest natural number.\nThis is one of the two constructors of `Nat`. \n***\n*import Init.Prelude*", + "The left hand side of an induction arm, `| foo a b c` or `| @foo a b c`\nwhere `foo` is a constructor of the inductive type and `a b c` are the arguments\nto the constructor.\n", "kind": "markdown"}} {"textDocument": {"uri": "file:///hover.lean"}, "position": {"line": 249, "character": 9}} @@ -588,7 +589,7 @@ {"range": {"start": {"line": 279, "character": 8}, "end": {"line": 279, "character": 10}}, - "contents": {"value": "```lean\n_e : 1 = Nat.zero\n```", "kind": "markdown"}} + "contents": {"value": "```lean\n_e : 1 = 0\n```", "kind": "markdown"}} {"textDocument": {"uri": "file:///hover.lean"}, "position": {"line": 282, "character": 9}} {"range": diff --git a/tests/lean/interactive/infoIssues.lean.expected.out b/tests/lean/interactive/infoIssues.lean.expected.out index 699aff477070..8a76f677174c 100644 --- a/tests/lean/interactive/infoIssues.lean.expected.out +++ b/tests/lean/interactive/infoIssues.lean.expected.out @@ -6,18 +6,18 @@ {"textDocument": {"uri": "file:///infoIssues.lean"}, "position": {"line": 13, "character": 33}} {"rendered": - "```lean\np : Prop\nn✝ : Nat\n⊢ Nat.succ n✝ = 0 + (0 + Nat.succ n✝)\n```\n---\n```lean\np : Prop\nn✝ : Nat\n⊢ Nat.succ n✝ = 0 + Nat.succ n✝\n```\n---\n```lean\np : Prop\nn✝ : Nat\n⊢ Nat.succ n✝ = Nat.succ n✝\n```", + "```lean\np : Prop\nn✝ : Nat\n⊢ n✝ + 1 = 0 + (0 + (n✝ + 1))\n```\n---\n```lean\np : Prop\nn✝ : Nat\n⊢ n✝ + 1 = 0 + (n✝ + 1)\n```\n---\n```lean\np : Prop\nn✝ : Nat\n⊢ n✝ + 1 = n✝ + 1\n```", "goals": - ["p : Prop\nn✝ : Nat\n⊢ Nat.succ n✝ = 0 + (0 + Nat.succ n✝)", - "p : Prop\nn✝ : Nat\n⊢ Nat.succ n✝ = 0 + Nat.succ n✝", - "p : Prop\nn✝ : Nat\n⊢ Nat.succ n✝ = Nat.succ n✝"]} + ["p : Prop\nn✝ : Nat\n⊢ n✝ + 1 = 0 + (0 + (n✝ + 1))", + "p : Prop\nn✝ : Nat\n⊢ n✝ + 1 = 0 + (n✝ + 1)", + "p : Prop\nn✝ : Nat\n⊢ n✝ + 1 = n✝ + 1"]} {"textDocument": {"uri": "file:///infoIssues.lean"}, "position": {"line": 13, "character": 36}} {"rendered": - "```lean\np : Prop\nn✝ : Nat\n⊢ Nat.succ n✝ = 0 + (0 + Nat.succ n✝)\n```\n---\n```lean\np : Prop\nn✝ : Nat\n⊢ Nat.succ n✝ = 0 + Nat.succ n✝\n```", + "```lean\np : Prop\nn✝ : Nat\n⊢ n✝ + 1 = 0 + (0 + (n✝ + 1))\n```\n---\n```lean\np : Prop\nn✝ : Nat\n⊢ n✝ + 1 = 0 + (n✝ + 1)\n```", "goals": - ["p : Prop\nn✝ : Nat\n⊢ Nat.succ n✝ = 0 + (0 + Nat.succ n✝)", - "p : Prop\nn✝ : Nat\n⊢ Nat.succ n✝ = 0 + Nat.succ n✝"]} + ["p : Prop\nn✝ : Nat\n⊢ n✝ + 1 = 0 + (0 + (n✝ + 1))", + "p : Prop\nn✝ : Nat\n⊢ n✝ + 1 = 0 + (n✝ + 1)"]} {"textDocument": {"uri": "file:///infoIssues.lean"}, "position": {"line": 15, "character": 2}} {"rendered": "```lean\ncase right\np : Prop\nx : Nat\n⊢ p\n```", diff --git a/tests/lean/interactive/plainGoal.lean.expected.out b/tests/lean/interactive/plainGoal.lean.expected.out index 40fa13fb654f..91a2fdd633bb 100644 --- a/tests/lean/interactive/plainGoal.lean.expected.out +++ b/tests/lean/interactive/plainGoal.lean.expected.out @@ -24,14 +24,13 @@ "goals": ["α : Sort ?u\n⊢ α → α"]} {"textDocument": {"uri": "file:///plainGoal.lean"}, "position": {"line": 15, "character": 9}} -{"rendered": "```lean\ncase zero\n⊢ 0 + Nat.zero = Nat.zero\n```", - "goals": ["case zero\n⊢ 0 + Nat.zero = Nat.zero"]} +{"rendered": "```lean\ncase zero\n⊢ 0 + 0 = 0\n```", + "goals": ["case zero\n⊢ 0 + 0 = 0"]} {"textDocument": {"uri": "file:///plainGoal.lean"}, "position": {"line": 17, "character": 5}} {"rendered": - "```lean\ncase succ\nn✝ : Nat\nn_ih✝ : 0 + n✝ = n✝\n⊢ 0 + Nat.succ n✝ = Nat.succ n✝\n```", - "goals": - ["case succ\nn✝ : Nat\nn_ih✝ : 0 + n✝ = n✝\n⊢ 0 + Nat.succ n✝ = Nat.succ n✝"]} + "```lean\ncase succ\nn✝ : Nat\na✝ : 0 + n✝ = n✝\n⊢ 0 + (n✝ + 1) = n✝ + 1\n```", + "goals": ["case succ\nn✝ : Nat\na✝ : 0 + n✝ = n✝\n⊢ 0 + (n✝ + 1) = n✝ + 1"]} {"textDocument": {"uri": "file:///plainGoal.lean"}, "position": {"line": 21, "character": 9}} {"rendered": "```lean\nα : Sort ?u\na : α\n⊢ α\n```", @@ -57,22 +56,21 @@ {"rendered": "no goals", "goals": []} {"textDocument": {"uri": "file:///plainGoal.lean"}, "position": {"line": 34, "character": 3}} -{"rendered": "```lean\ncase zero\n⊢ 0 + Nat.zero = Nat.zero\n```", - "goals": ["case zero\n⊢ 0 + Nat.zero = Nat.zero"]} +{"rendered": "```lean\ncase zero\n⊢ 0 + 0 = 0\n```", + "goals": ["case zero\n⊢ 0 + 0 = 0"]} {"textDocument": {"uri": "file:///plainGoal.lean"}, "position": {"line": 40, "character": 3}} {"rendered": - "```lean\ncase zero\n⊢ 0 + Nat.zero = Nat.zero\n```\n---\n```lean\ncase succ\nn✝ : Nat\nn_ih✝ : 0 + n✝ = n✝\n⊢ 0 + Nat.succ n✝ = Nat.succ n✝\n```", + "```lean\ncase zero\n⊢ 0 + 0 = 0\n```\n---\n```lean\ncase succ\nn✝ : Nat\na✝ : 0 + n✝ = n✝\n⊢ 0 + (n✝ + 1) = n✝ + 1\n```", "goals": - ["case zero\n⊢ 0 + Nat.zero = Nat.zero", - "case succ\nn✝ : Nat\nn_ih✝ : 0 + n✝ = n✝\n⊢ 0 + Nat.succ n✝ = Nat.succ n✝"]} + ["case zero\n⊢ 0 + 0 = 0", + "case succ\nn✝ : Nat\na✝ : 0 + n✝ = n✝\n⊢ 0 + (n✝ + 1) = n✝ + 1"]} {"textDocument": {"uri": "file:///plainGoal.lean"}, "position": {"line": 44, "character": 3}} {"rendered": - "```lean\ncase zero\n⊢ 0 + Nat.zero = Nat.zero\n```\n---\n```lean\ncase succ\nn✝ : Nat\n⊢ 0 + Nat.succ n✝ = Nat.succ n✝\n```", + "```lean\ncase zero\n⊢ 0 + 0 = 0\n```\n---\n```lean\ncase succ\nn✝ : Nat\n⊢ 0 + (n✝ + 1) = n✝ + 1\n```", "goals": - ["case zero\n⊢ 0 + Nat.zero = Nat.zero", - "case succ\nn✝ : Nat\n⊢ 0 + Nat.succ n✝ = Nat.succ n✝"]} + ["case zero\n⊢ 0 + 0 = 0", "case succ\nn✝ : Nat\n⊢ 0 + (n✝ + 1) = n✝ + 1"]} {"textDocument": {"uri": "file:///plainGoal.lean"}, "position": {"line": 48, "character": 3}} {"rendered": "```lean\na b : Nat\n⊢ a = b\n```", @@ -104,11 +102,11 @@ {"textDocument": {"uri": "file:///plainGoal.lean"}, "position": {"line": 78, "character": 29}} {"rendered": - "```lean\nt a n✝ : Nat\nn_ih✝ : t * (a + n✝) = t * a + t * n✝\n⊢ t * (a + n✝) + t = t * a + t * Nat.succ n✝\n```\n---\n```lean\nt a n✝ : Nat\nn_ih✝ : t * (a + n✝) = t * a + t * n✝\n⊢ t * (a + n✝) + t = t * a + (t * n✝ + t)\n```\n---\n```lean\nt a n✝ : Nat\nn_ih✝ : t * (a + n✝) = t * a + t * n✝\n⊢ t * (a + n✝) + t = t * a + (t * n✝ + t)\n```", + "```lean\nt a n✝ : Nat\na✝ : t * (a + n✝) = t * a + t * n✝\n⊢ t * (a + n✝) + t = t * a + t * (n✝ + 1)\n```\n---\n```lean\nt a n✝ : Nat\na✝ : t * (a + n✝) = t * a + t * n✝\n⊢ t * (a + n✝) + t = t * a + (t * n✝ + t)\n```\n---\n```lean\nt a n✝ : Nat\na✝ : t * (a + n✝) = t * a + t * n✝\n⊢ t * (a + n✝) + t = t * a + (t * n✝ + t)\n```", "goals": - ["t a n✝ : Nat\nn_ih✝ : t * (a + n✝) = t * a + t * n✝\n⊢ t * (a + n✝) + t = t * a + t * Nat.succ n✝", - "t a n✝ : Nat\nn_ih✝ : t * (a + n✝) = t * a + t * n✝\n⊢ t * (a + n✝) + t = t * a + (t * n✝ + t)", - "t a n✝ : Nat\nn_ih✝ : t * (a + n✝) = t * a + t * n✝\n⊢ t * (a + n✝) + t = t * a + (t * n✝ + t)"]} + ["t a n✝ : Nat\na✝ : t * (a + n✝) = t * a + t * n✝\n⊢ t * (a + n✝) + t = t * a + t * (n✝ + 1)", + "t a n✝ : Nat\na✝ : t * (a + n✝) = t * a + t * n✝\n⊢ t * (a + n✝) + t = t * a + (t * n✝ + t)", + "t a n✝ : Nat\na✝ : t * (a + n✝) = t * a + t * n✝\n⊢ t * (a + n✝) + t = t * a + (t * n✝ + t)"]} {"textDocument": {"uri": "file:///plainGoal.lean"}, "position": {"line": 82, "character": 53}} {"rendered": diff --git a/tests/lean/run/1337.lean b/tests/lean/run/1337.lean index aa914ff820d5..8e6a2a2068fe 100644 --- a/tests/lean/run/1337.lean +++ b/tests/lean/run/1337.lean @@ -3,7 +3,7 @@ theorem n_minus_one_le_n {n : Nat} : n > 0 → n - 1 < n := by | zero => simp [] | succ n => intros - rw [Nat.succ_eq_add_one, Nat.add_sub_cancel] + rw [Nat.add_sub_cancel] apply Nat.le.refl partial def foo : Array Int → Int diff --git a/tests/lean/run/aStructPerfIssue.lean b/tests/lean/run/aStructPerfIssue.lean index 2e85d2b8f1de..742f2fcee1c4 100644 --- a/tests/lean/run/aStructPerfIssue.lean +++ b/tests/lean/run/aStructPerfIssue.lean @@ -6,7 +6,7 @@ universe u v w inductive Id {A : Type u} : A → A → Type u | refl {a : A} : Id a a -attribute [eliminator] Id.casesOn +attribute [cases_eliminator] Id.casesOn infix:50 (priority := high) " = " => Id @@ -40,7 +40,7 @@ def Iff.comp {A : Type u} {B : Type v} {C : Type w} : λ p q => (q.left ∘ p.left, p.right ∘ q.right) inductive Empty : Type u -attribute [eliminator] Empty.casesOn +attribute [cases_eliminator] Empty.casesOn notation "𝟎" => Empty @@ -88,7 +88,7 @@ notation n "-Type" => nType n inductive Unit : Type u | star : Unit -attribute [eliminator] Unit.casesOn +attribute [cases_eliminator] Unit.casesOn def Homotopy {A : Type u} {B : A → Type v} (f g : ∀ x, B x) := ∀ (x : A), f x = g x diff --git a/tests/lean/run/casesAnyTypeIssue.lean b/tests/lean/run/casesAnyTypeIssue.lean index c0cf4d524762..8e76165419a1 100644 --- a/tests/lean/run/casesAnyTypeIssue.lean +++ b/tests/lean/run/casesAnyTypeIssue.lean @@ -3,7 +3,7 @@ namespace MWE inductive Id {A : Type u} : A → A → Type u | refl {a : A} : Id a a -attribute [eliminator] Id.casesOn +attribute [induction_eliminator] Id.casesOn infix:50 (priority := high) " = " => Id diff --git a/tests/lean/run/compatibleTypesEtaIssue.lean b/tests/lean/run/compatibleTypesEtaIssue.lean index bc19cb9d347c..9257b24bb313 100644 --- a/tests/lean/run/compatibleTypesEtaIssue.lean +++ b/tests/lean/run/compatibleTypesEtaIssue.lean @@ -5,8 +5,6 @@ universe u v w inductive Id {A : Type u} : A → A → Type u | refl {a : A} : Id a a -attribute [eliminator] Id.casesOn - infix:50 (priority := high) " = " => Id def contr (A : Type u) := Σ (a : A), ∀ b, a = b diff --git a/tests/lean/run/defaultEliminator.lean b/tests/lean/run/defaultEliminator.lean index 6a29dcebd43e..588487931503 100644 --- a/tests/lean/run/defaultEliminator.lean +++ b/tests/lean/run/defaultEliminator.lean @@ -1,4 +1,4 @@ -@[eliminator] protected def Nat.recDiag {motive : Nat → Nat → Sort u} +@[induction_eliminator] protected def Nat.recDiag {motive : Nat → Nat → Sort u} (zero_zero : motive 0 0) (succ_zero : (x : Nat) → motive x 0 → motive (x + 1) 0) (zero_succ : (y : Nat) → motive 0 y → motive 0 (y + 1)) diff --git a/tests/lean/run/eliminatorImplicitTargets.lean b/tests/lean/run/eliminatorImplicitTargets.lean index 83548111f87d..c0bc93db170c 100644 --- a/tests/lean/run/eliminatorImplicitTargets.lean +++ b/tests/lean/run/eliminatorImplicitTargets.lean @@ -3,7 +3,7 @@ inductive Equality {α : Type u} : α → α → Type u open Equality -@[eliminator] +@[induction_eliminator] def ind {α : Type u} (motive : ∀ (a b : α) (p : Equality a b), Sort v) {a : α} (πrefl : motive a a refl) {b : α} (p : Equality a b) : motive a b p := @Equality.casesOn α a (λ b p => motive a a refl → motive a b p) b p diff --git a/tests/lean/run/internalizeCasesIssue.lean b/tests/lean/run/internalizeCasesIssue.lean index 8233bde6a05e..d7066fa71945 100644 --- a/tests/lean/run/internalizeCasesIssue.lean +++ b/tests/lean/run/internalizeCasesIssue.lean @@ -3,7 +3,7 @@ namespace MWE inductive Id {A : Type u} : A → A → Type u | refl {a : A} : Id a a -attribute [eliminator] Id.casesOn +attribute [induction_eliminator] Id.casesOn infix:50 (priority := high) " = " => Id diff --git a/tests/lean/run/lcnfCastIssue.lean b/tests/lean/run/lcnfCastIssue.lean index cdef50c3b1d0..4a80a1a13233 100644 --- a/tests/lean/run/lcnfCastIssue.lean +++ b/tests/lean/run/lcnfCastIssue.lean @@ -5,14 +5,14 @@ universe u v w inductive Id {A : Type u} : A → A → Type u | refl {a : A} : Id a a -attribute [eliminator] Id.casesOn +attribute [induction_eliminator] Id.casesOn infix:50 (priority := high) " = " => Id inductive Unit : Type u | star : Unit -attribute [eliminator] Unit.casesOn +attribute [induction_eliminator] Unit.casesOn notation "𝟏" => Unit notation "★" => Unit.star diff --git a/tests/lean/run/lcnfInliningIssue.lean b/tests/lean/run/lcnfInliningIssue.lean index bcceecefd093..26ed29205fea 100644 --- a/tests/lean/run/lcnfInliningIssue.lean +++ b/tests/lean/run/lcnfInliningIssue.lean @@ -3,7 +3,7 @@ namespace MWE inductive Id {A : Type u} : A → A → Type u | refl {a : A} : Id a a -attribute [eliminator] Id.casesOn +attribute [induction_eliminator] Id.casesOn infix:50 (priority := high) " = " => Id diff --git a/tests/lean/simp_trace.lean.expected.out b/tests/lean/simp_trace.lean.expected.out index 884be2f7ff02..081db169c2ab 100644 --- a/tests/lean/simp_trace.lean.expected.out +++ b/tests/lean/simp_trace.lean.expected.out @@ -6,7 +6,7 @@ Try this: simp only [length, gt_iff_lt] [Meta.Tactic.simp.rewrite] unfold length, length (b :: as) ==> length as + 1 [Meta.Tactic.simp.rewrite] @gt_iff_lt:1000, length as + 1 + 1 > length as ==> length as < length as + 1 + 1 Try this: simp only [fact, gt_iff_lt, Nat.zero_lt_succ, Nat.mul_pos_iff_of_pos_left] -[Meta.Tactic.simp.rewrite] unfold fact, fact (Nat.succ x) ==> (x + 1) * fact x +[Meta.Tactic.simp.rewrite] unfold fact, fact (x + 1) ==> (x + 1) * fact x [Meta.Tactic.simp.rewrite] @gt_iff_lt:1000, (x + 1) * fact x > 0 ==> 0 < (x + 1) * fact x [Meta.Tactic.simp.rewrite] Nat.zero_lt_succ:1000, 0 < x + 1 ==> True [Meta.Tactic.simp.rewrite] @Nat.mul_pos_iff_of_pos_left:1000, 0 < (x + 1) * fact x ==> 0 < fact x diff --git a/tests/lean/unfold1.lean.expected.out b/tests/lean/unfold1.lean.expected.out index 88c8a3cb05d6..d046ef4aa81a 100644 --- a/tests/lean/unfold1.lean.expected.out +++ b/tests/lean/unfold1.lean.expected.out @@ -2,7 +2,7 @@ unfold1.lean:22:4-22:8: error: simp made no progress case succ x : Nat ih : isEven (2 * x) = true -⊢ (match 2 * Nat.succ x with +⊢ (match 2 * (x + 1) with | 0 => true | Nat.succ n => isOdd n) = true diff --git a/tests/lean/unsolvedIndCases.lean.expected.out b/tests/lean/unsolvedIndCases.lean.expected.out index bd3736bee15f..7c6ce644e194 100644 --- a/tests/lean/unsolvedIndCases.lean.expected.out +++ b/tests/lean/unsolvedIndCases.lean.expected.out @@ -1,22 +1,22 @@ unsolvedIndCases.lean:3:11-3:18: error: unsolved goals case zero -⊢ 0 + Nat.zero = Nat.zero +⊢ 0 + 0 = 0 unsolvedIndCases.lean:4:11-4:18: error: unsolved goals case succ y : Nat -⊢ 0 + Nat.succ y = Nat.succ y +⊢ 0 + (y + 1) = y + 1 unsolvedIndCases.lean:8:14-8:21: error: unsolved goals case zero -⊢ 0 + Nat.zero = Nat.zero +⊢ 0 + 0 = 0 unsolvedIndCases.lean:9:14-9:21: error: unsolved goals case succ y : Nat ih : 0 + y = y -⊢ 0 + Nat.succ y = Nat.succ y +⊢ 0 + (y + 1) = y + 1 unsolvedIndCases.lean:14:11-14:18: error: unsolved goals case succ y : Nat -⊢ 0 + Nat.succ y = Nat.succ y +⊢ 0 + (y + 1) = y + 1 unsolvedIndCases.lean:18:18-18:25: error: unsolved goals case ind x y : Nat From 1d3ef577c28b6298f4edf8aef86b23ca3ba99b66 Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Sat, 9 Mar 2024 15:48:41 -0800 Subject: [PATCH 09/32] chore: disable some tests on Windows (#3642) This is a temporary workaround for a limitation on Windows shared libraries. We are getting errors of the form: ``` ld.lld: error: too many exported symbols (got 65572, max 65535) ``` --- src/shell/CMakeLists.txt | 42 ++++++++++++++++++++++++---------------- 1 file changed, 25 insertions(+), 17 deletions(-) diff --git a/src/shell/CMakeLists.txt b/src/shell/CMakeLists.txt index 87da1764b43b..9bcc910ad121 100644 --- a/src/shell/CMakeLists.txt +++ b/src/shell/CMakeLists.txt @@ -119,14 +119,18 @@ ENDFOREACH(T) # LEAN BENCHMARK TESTS # do not test all .lean files in bench/ -file(GLOB LEANBENCHTESTS "${LEAN_SOURCE_DIR}/../tests/bench/*.lean.expected.out") -FOREACH(T_OUT ${LEANBENCHTESTS}) - string(REPLACE ".expected.out" "" T ${T_OUT}) - GET_FILENAME_COMPONENT(T_NAME ${T} NAME) - add_test(NAME "leanbenchtest_${T_NAME}" - WORKING_DIRECTORY "${LEAN_SOURCE_DIR}/../tests/bench" - COMMAND bash -c "${TEST_VARS} ./test_single.sh ${T_NAME}") -ENDFOREACH(T_OUT) +if(${CMAKE_SYSTEM_NAME} MATCHES "Windows") + message(STATUS "Skipping compiler tests on Windows because of shared library limit on number of exported symbols") +else() + file(GLOB LEANBENCHTESTS "${LEAN_SOURCE_DIR}/../tests/bench/*.lean.expected.out") + FOREACH(T_OUT ${LEANBENCHTESTS}) + string(REPLACE ".expected.out" "" T ${T_OUT}) + GET_FILENAME_COMPONENT(T_NAME ${T} NAME) + add_test(NAME "leanbenchtest_${T_NAME}" + WORKING_DIRECTORY "${LEAN_SOURCE_DIR}/../tests/bench" + COMMAND bash -c "${TEST_VARS} ./test_single.sh ${T_NAME}") + ENDFOREACH(T_OUT) +endif() file(GLOB LEANINTERPTESTS "${LEAN_SOURCE_DIR}/../tests/plugin/*.lean") FOREACH(T ${LEANINTERPTESTS}) @@ -146,15 +150,19 @@ FOREACH(T ${LEANT0TESTS}) ENDFOREACH(T) # LEAN PACKAGE TESTS -file(GLOB LEANPKGTESTS "${LEAN_SOURCE_DIR}/../tests/pkg/*") -FOREACH(T ${LEANPKGTESTS}) - if(IS_DIRECTORY ${T}) - GET_FILENAME_COMPONENT(T_NAME ${T} NAME) - add_test(NAME "leanpkgtest_${T_NAME}" - WORKING_DIRECTORY "${T}" - COMMAND bash -c "${TEST_VARS} ./test.sh") - endif() -ENDFOREACH(T) +if(${CMAKE_SYSTEM_NAME} MATCHES "Windows") + message(STATUS "Skipping compiler tests on Windows because of shared library limit on number of exported symbols") +else() + file(GLOB LEANPKGTESTS "${LEAN_SOURCE_DIR}/../tests/pkg/*") + FOREACH(T ${LEANPKGTESTS}) + if(IS_DIRECTORY ${T}) + GET_FILENAME_COMPONENT(T_NAME ${T} NAME) + add_test(NAME "leanpkgtest_${T_NAME}" + WORKING_DIRECTORY "${T}" + COMMAND bash -c "${TEST_VARS} ./test.sh") + endif() + ENDFOREACH(T) +endif() # LEAN SERVER TESTS file(GLOB LEANTESTS "${LEAN_SOURCE_DIR}/../tests/lean/server/*.lean") From 32dcc6eb895b58df3d3241a2521963e64995b621 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sun, 10 Mar 2024 23:57:10 +0100 Subject: [PATCH 10/32] feat: GuessLex: avoid writing `sizeOf` in termination argument when not needed (#3630) this makes `termination_by?` even slicker. The heuristics is agressive in the non-mutual case (will omit `sizeOf` if the argument is non-dependent and the `WellFoundedRelation` relation is via `sizeOfWFRel`. In the mutual case we'd also have to check the arguments, as they line up in the termination argument, have the same types. I did not bother at this point; in the mutual case we omit `sizeOf` only if the argument type is `Nat`. As a drive-by fix, `termination_by?` now also works on functions that have only one plausible measure. --- src/Lean/Elab/PreDefinition/WF/GuessLex.lean | 119 +++++++++++++----- tests/lean/guessLex.lean | 98 +++++++++++++-- tests/lean/guessLex.lean.expected.out | 56 ++++++--- tests/lean/guessLexTricky.lean.expected.out | 6 +- tests/lean/guessLexTricky2.lean.expected.out | 4 +- .../interactive/terminationBySuggestion.lean | 14 ++- .../terminationBySuggestion.lean.expected.out | 32 ++++- tests/lean/issue2981.lean.expected.out | 4 +- 8 files changed, 264 insertions(+), 69 deletions(-) diff --git a/src/Lean/Elab/PreDefinition/WF/GuessLex.lean b/src/Lean/Elab/PreDefinition/WF/GuessLex.lean index 419f6b8ed0af..d20eebfab03d 100644 --- a/src/Lean/Elab/PreDefinition/WF/GuessLex.lean +++ b/src/Lean/Elab/PreDefinition/WF/GuessLex.lean @@ -302,7 +302,11 @@ def GuessLexRel.toNatRel : GuessLexRel → Expr | le => mkAppN (mkConst ``LE.le [levelZero]) #[mkConst ``Nat, mkConst ``instLENat] | no_idea => unreachable! -/-- Given an expression `e`, produce `sizeOf e` with a suitable instance. -/ +/-- +Given an expression `e`, produce `sizeOf e` with a suitable instance. +NB: We must use the instance of the type of the function parameter! +The concrete argument at hand may have a different (still def-eq) typ. +-/ def mkSizeOf (e : Expr) : MetaM Expr := do let ty ← inferType e let lvl ← getLevel ty @@ -315,8 +319,8 @@ def mkSizeOf (e : Expr) : MetaM Expr := do For a given recursive call, and a choice of parameter and argument index, try to prove equality, < or ≤. -/ -def evalRecCall (decrTactic? : Option DecreasingBy) (rcc : RecCallWithContext) (paramIdx argIdx : Nat) : - MetaM GuessLexRel := do +def evalRecCall (decrTactic? : Option DecreasingBy) (rcc : RecCallWithContext) + (paramIdx argIdx : Nat) : MetaM GuessLexRel := do rcc.ctxt.run do let param := rcc.params[paramIdx]! let arg := rcc.args[argIdx]! @@ -407,7 +411,7 @@ def inspectCall (rc : RecCallCache) : MutualMeasure → MetaM GuessLexRel return .eq /-- -Given a predefinition with value `fun (x_₁ ... xₙ) (y_₁ : α₁)... (yₘ : αₘ) => ...`, +Given a predefinition with value `fun (x₁ ... xₙ) (y₁ : α₁)... (yₘ : αₘ) => ...`, where `n = fixedPrefixSize`, return an array `A` s.t. `i ∈ A` iff `sizeOf yᵢ` reduces to a literal. This is the case for types such as `Prop`, `Type u`, etc. These arguments should not be considered when guessing a well-founded relation. @@ -425,6 +429,47 @@ def getForbiddenByTrivialSizeOf (fixedPrefixSize : Nat) (preDef : PreDefinition) result := result.push i return result +/-- +Given a predefinition with value `fun (x₁ ... xₙ) (y₁ : α₁)... (yₘ : αₘ) => ...`, +where `n = fixedPrefixSize`, return an array `A` s.t. `i ∈ A` iff the +`WellFoundedRelation` of `aᵢ` goes via `SizeOf`, and `aᵢ` does not depend on `y₁`…. + +These are the parameters for which we omit an explicit call to `sizeOf` in the termination argument. + +We only use this in the non-mutual case; in the mutual case we would have to additional check +if the parameters that line up in the actual `TerminationWF` have the same type. +-/ +def getSizeOfParams (fixedPrefixSize : Nat) (preDef : PreDefinition) : MetaM (Array Nat) := + lambdaTelescope preDef.value fun xs _ => do + let xs : Array Expr := xs[fixedPrefixSize:] + let mut result := #[] + for x in xs, i in [:xs.size] do + try + let t ← inferType x + if t.hasAnyFVar (fun fvar => xs.contains (.fvar fvar)) then continue + let u ← getLevel t + let wfi ← synthInstance (.app (.const ``WellFoundedRelation [u]) t) + let soi ← synthInstance (.app (.const ``SizeOf [u]) t) + if ← isDefEq wfi (mkApp2 (.const ``sizeOfWFRel [u]) t soi) then + result := result.push i + catch _ => + pure () + return result + +/-- +Given a predefinition with value `fun (x₁ ... xₙ) (y₁ : α₁)... (yₘ : αₘ) => ...`, +where `n = fixedPrefixSize`, return an array `A` s.t. `i ∈ A` iff `aᵢ` is `Nat`. +These are parameters where we can definitely omit the call to `sizeOf`. +-/ +def getNatParams (fixedPrefixSize : Nat) (preDef : PreDefinition) : MetaM (Array Nat) := + lambdaTelescope preDef.value fun xs _ => do + let xs : Array Expr := xs[fixedPrefixSize:] + let mut result := #[] + for x in xs, i in [:xs.size] do + let t ← inferType x + if ← withReducible (isDefEq t (.const `Nat [])) then + result := result.push i + return result /-- Generate all combination of arguments, skipping those that are forbidden. @@ -539,23 +584,26 @@ combination of these measures. The parameters are * `measures`: The measures to be used. -/ def buildTermWF (originalVarNamess : Array (Array Name)) (varNamess : Array (Array Name)) - (measures : Array MutualMeasure) : MetaM TerminationWF := do + (needsNoSizeOf : Array (Array Nat)) (measures : Array MutualMeasure) : MetaM TerminationWF := do varNamess.mapIdxM fun funIdx varNames => do let idents := varNames.map mkIdent let measureStxs ← measures.mapM fun | .args varIdxs => do let varIdx := varIdxs[funIdx]! let v := idents[varIdx]! - -- Print `sizeOf` as such, unless it is shadowed. - -- Shadowing by a `def` in the current namespace is handled by `unresolveNameGlobal`. - -- But it could also be shadowed by an earlier parameter (including the fixed prefix), - -- so look for unqualified (single tick) occurrences in `originalVarNames` - let sizeOfIdent := - if originalVarNamess[funIdx]!.any (· = `sizeOf) then - mkIdent ``sizeOf -- fully qualified - else - mkIdent (← unresolveNameGlobal ``sizeOf) - `($sizeOfIdent $v) + if needsNoSizeOf[funIdx]!.contains varIdx then + `($v) + else + -- Print `sizeOf` as such, unless it is shadowed. + -- Shadowing by a `def` in the current namespace is handled by `unresolveNameGlobal`. + -- But it could also be shadowed by an earlier parameter (including the fixed prefix), + -- so look for unqualified (single tick) occurrences in `originalVarNames` + let sizeOfIdent := + if originalVarNamess[funIdx]!.any (· = `sizeOf) then + mkIdent ``sizeOf -- fully qualified + else + mkIdent (← unresolveNameGlobal ``sizeOf) + `($sizeOfIdent $v) | .func funIdx' => if funIdx' == funIdx then `(1) else `(0) let body ← mkTupleSyntax measureStxs return { ref := .missing, vars := idents, body, synthetic := true } @@ -668,11 +716,20 @@ def explainFailure (declNames : Array Name) (varNamess : Array (Array Name)) r := r ++ (← explainMutualFailure declNames varNamess rcs) return r -end Lean.Elab.WF.GuessLex - -namespace Lean.Elab.WF +/-- +Shows the termination measure used to the user, and implements `termination_by?` +-/ +def reportWF (preDefs : Array PreDefinition) (wf : TerminationWF) : MetaM Unit := do + let extraParamss := preDefs.map (·.termination.extraParams) + let wf' := trimTermWF extraParamss wf + for preDef in preDefs, term in wf' do + if showInferredTerminationBy.get (← getOptions) then + logInfoAt preDef.ref m!"Inferred termination argument:\n{← term.unexpand}" + if let some ref := preDef.termination.terminationBy?? then + Tactic.TryThis.addSuggestion ref (← term.unexpand) -open Lean.Elab.WF.GuessLex +end GuessLex +open GuessLex /-- Main entry point of this module: @@ -683,14 +740,17 @@ terminates. See the module doc string for a high-level overview. def guessLex (preDefs : Array PreDefinition) (unaryPreDef : PreDefinition) (fixedPrefixSize : Nat) : MetaM TerminationWF := do - let extraParamss := preDefs.map (·.termination.extraParams) let originalVarNamess ← preDefs.mapM originalVarNames let varNamess ← originalVarNamess.mapM (naryVarNames fixedPrefixSize ·) let arities := varNamess.map (·.size) trace[Elab.definition.wf] "varNames is: {varNamess}" - let forbiddenArgs ← preDefs.mapM fun preDef => - getForbiddenByTrivialSizeOf fixedPrefixSize preDef + let forbiddenArgs ← preDefs.mapM (getForbiddenByTrivialSizeOf fixedPrefixSize) + let needsNoSizeOf ← + if preDefs.size = 1 then + preDefs.mapM (getSizeOfParams fixedPrefixSize) + else + preDefs.mapM (getNatParams fixedPrefixSize) -- The list of measures, including the measures that order functions. -- The function ordering measures come last @@ -698,7 +758,9 @@ def guessLex (preDefs : Array PreDefinition) (unaryPreDef : PreDefinition) -- If there is only one plausible measure, use that if let #[solution] := measures then - return ← buildTermWF originalVarNamess varNamess #[solution] + let wf ← buildTermWF originalVarNamess varNamess needsNoSizeOf #[solution] + reportWF preDefs wf + return wf -- Collect all recursive calls and extract their context let recCalls ← collectRecCalls unaryPreDef fixedPrefixSize arities @@ -708,15 +770,8 @@ def guessLex (preDefs : Array PreDefinition) (unaryPreDef : PreDefinition) match ← liftMetaM <| solve measures callMatrix with | .some solution => do - let wf ← buildTermWF originalVarNamess varNamess solution - - let wf' := trimTermWF extraParamss wf - for preDef in preDefs, term in wf' do - if showInferredTerminationBy.get (← getOptions) then - logInfoAt preDef.ref m!"Inferred termination argument:\n{← term.unexpand}" - if let some ref := preDef.termination.terminationBy?? then - Tactic.TryThis.addSuggestion ref (← term.unexpand) - + let wf ← buildTermWF originalVarNamess varNamess needsNoSizeOf solution + reportWF preDefs wf return wf | .none => let explanation ← explainFailure (preDefs.map (·.declName)) varNamess rcs diff --git a/tests/lean/guessLex.lean b/tests/lean/guessLex.lean index 9fabe2b711a3..c836ab72d75e 100644 --- a/tests/lean/guessLex.lean +++ b/tests/lean/guessLex.lean @@ -87,6 +87,7 @@ def confuseLex2 : @PSigma Nat (fun _ => Nat) → Nat | ⟨0,_⟩ => 0 | ⟨.succ y,.succ n⟩ => confuseLex2 ⟨y,n⟩ +-- NB: uses sizeOf to make the termination argument non-dependent def dependent : (n : Nat) → (m : Fin n) → Nat | 0, i => Fin.elim0 i | .succ 0, 0 => 0 @@ -94,6 +95,11 @@ def dependent : (n : Nat) → (m : Fin n) → Nat | .succ (.succ n), ⟨.succ m, h⟩ => dependent (.succ (.succ n)) ⟨m, Nat.lt_of_le_of_lt (Nat.le_succ _) h⟩ +-- NB: does not use sizeOf, as parameters in the fixed prefix are fine. +def dependentWithFixedPrefix (n : Nat) : (m : Fin n) → (acc : Nat) → Nat + | ⟨0, _⟩, acc => acc + | ⟨i+1, h⟩, acc => dependentWithFixedPrefix n ⟨i, Nat.lt_of_succ_lt h⟩ (acc + i) + -- An example based on a real world problem, condensed by Leo inductive Expr where | add (a b : Expr) @@ -110,6 +116,7 @@ def eval_add (a : Expr × Expr) : Nat := | (x, y) => eval x + eval y end + namespace VarNames /-! Test that varnames are inferred nicely. -/ @@ -127,24 +134,97 @@ def shadow2 (some_n : Nat) : Nat → Nat | .succ n => shadow2 (some_n + 1) n decreasing_by decreasing_tactic + +-- The following test whether `sizeOf` is properly printed, and possibly qualified +-- For this we need a type that needs an explicit “sizeOf”. + +structure OddNat where nat : Nat +instance : WellFoundedRelation OddNat := measure (fun ⟨n⟩ => n+1) + +-- Just to check that sizeOf is actually used +def oddNat : OddNat → Nat + | ⟨0⟩ => 0 + | ⟨.succ n⟩ => oddNat ⟨n⟩ +decreasing_by decreasing_tactic + -- Shadowing `sizeOf`, as a varying paramter -def shadowSizeOf1 (sizeOf : Nat) : Nat → Nat - | 0 => 0 - | .succ n => shadowSizeOf1 (sizeOf + 1) n +def shadowSizeOf1 (sizeOf : Nat) : OddNat → Nat + | ⟨0⟩ => 0 + | ⟨.succ n⟩ => shadowSizeOf1 (sizeOf + 1) ⟨n⟩ decreasing_by decreasing_tactic -- Shadowing `sizeOf`, as a fixed paramter -def shadowSizeOf2 (sizeOf : Nat) : Nat → Nat → Nat - | 0, m => m - | .succ n, m => shadowSizeOf2 sizeOf n m +def shadowSizeOf2 (sizeOf : Nat) : OddNat → Nat → Nat + | ⟨0⟩, m => m + | ⟨.succ n⟩, m => shadowSizeOf2 sizeOf ⟨n⟩ m decreasing_by decreasing_tactic -- Shadowing `sizeOf`, as something in the environment def sizeOf : Nat := 2 -def qualifiedSizeOf (m : Nat) : Nat → Nat - | 0 => 0 - | .succ n => qualifiedSizeOf (m + 1) n +def qualifiedSizeOf (m : Nat) : OddNat → Nat + | ⟨0⟩ => 0 + | ⟨.succ n⟩ => qualifiedSizeOf (m + 1) ⟨n⟩ decreasing_by decreasing_tactic end VarNames + + +namespace MutualNotNat1 + +-- A type that isn't Nat, checking that the inferred argument uses `sizeOf` so that +-- the types of the termination argument aligns. +structure OddNat2 where nat : Nat +instance : SizeOf OddNat2 := ⟨fun n => n.nat⟩ +@[simp] theorem OddNat2.sizeOf_eq (n : OddNat2) : sizeOf n = n.nat := rfl +mutual +def foo : Nat → Nat + | 0 => 0 + | n+1 => bar ⟨n⟩ +def bar : OddNat2 → Nat + | ⟨0⟩ => 0 + | ⟨n+1⟩ => foo n +end +end MutualNotNat1 + +namespace MutualNotNat2 +-- A type that is defeq to Nat, but with a different `sizeOf`, checking that the +-- inferred argument uses `sizeOf` so that the types of the termination argument aligns. +def OddNat3 := Nat +instance : SizeOf OddNat3 := ⟨fun n => 42 - @id Nat n⟩ +@[simp] theorem OddNat3.sizeOf_eq (n : OddNat3) : sizeOf n = 42 - @id Nat n := rfl +mutual +def foo : Nat → Nat + | 0 => 0 + | n+1 => + if h : n < 42 then bar (42 - n) else 0 + -- termination_by x1 => x1 + decreasing_by simp_wf; simp [OddNat3]; omega +def bar (o : OddNat3) : Nat := if h : @id Nat o < 41 then foo (41 - @id Nat o) else 0 + -- termination_by sizeOf o + decreasing_by simp_wf; simp [id] at *; omega +end +namespace MutualNotNat2 + +namespace MutualNotNat3 +-- A varant of the above, but where the type of the parameter refined to `Nat`. +-- This tests if `GuessLex` is inferring the `SizeOf` instance based on the type of the +-- concrete parameter/argument (wrong, but status quo), or based on the types in the function +-- signature (correct, todo) +def OddNat3 := Nat +instance : SizeOf OddNat3 := ⟨fun n => 42 - @id Nat n⟩ +@[simp] theorem OddNat3.sizeOf_eq (n : OddNat3) : sizeOf n = 42 - @id Nat n := rfl +mutual +def foo : Nat → Nat + | 0 => 0 + | n+1 => + if h : n < 42 then bar (42 - n) else 0 + -- termination_by x1 => x1 + decreasing_by simp_wf; simp [OddNat3]; omega +def bar : OddNat3 → Nat + | Nat.zero => 0 + | n+1 => if h : n < 41 then foo (40 - n) else 0 + -- termination_by x1 => sizeOf x1 + decreasing_by simp_wf; omega +end +namespace MutualNotNat3 diff --git a/tests/lean/guessLex.lean.expected.out b/tests/lean/guessLex.lean.expected.out index 7101d313ae0c..737e10af6a4e 100644 --- a/tests/lean/guessLex.lean.expected.out +++ b/tests/lean/guessLex.lean.expected.out @@ -1,41 +1,65 @@ Inferred termination argument: -termination_by (sizeOf n, sizeOf m) +termination_by (n, m) Inferred termination argument: -termination_by (sizeOf m, sizeOf n) +termination_by (m, n) Inferred termination argument: -termination_by (sizeOf n, sizeOf m) +termination_by (n, m) Inferred termination argument: -termination_by x1 x2 => (sizeOf x2, sizeOf x1) +termination_by x1 x2 => (x2, x1) Inferred termination argument: -termination_by x1 => sizeOf x1 +termination_by x1 => x1 Inferred termination argument: -termination_by x1 => sizeOf x1 +termination_by x1 => x1 Inferred termination argument: -termination_by x1 => sizeOf x1 +termination_by x1 => x1 Inferred termination argument: -termination_by x1 => sizeOf x1 +termination_by x1 => x1 Inferred termination argument: -termination_by x1 => (sizeOf x1, 0) +termination_by x1 => (x1, 0) Inferred termination argument: -termination_by (sizeOf n, 1) +termination_by (n, 1) Inferred termination argument: -termination_by (sizeOf m, sizeOf n) +termination_by (m, n) +Inferred termination argument: +termination_by x1 x2 x3 x4 x5 x6 x7 x8 => (x8, x7, x6, x5, x4, x3, x2, x1) +Inferred termination argument: +termination_by x1 => sizeOf x1 Inferred termination argument: -termination_by x1 x2 x3 x4 x5 x6 x7 x8 => - (sizeOf x8, sizeOf x7, sizeOf x6, sizeOf x5, sizeOf x4, sizeOf x3, sizeOf x2, sizeOf x1) +termination_by x1 x2 => (x1, sizeOf x2) Inferred termination argument: -termination_by x1 x2 => (sizeOf x1, sizeOf x2) +termination_by x1 x2 => x1 Inferred termination argument: termination_by (sizeOf a, 1) Inferred termination argument: termination_by (sizeOf a, 0) Inferred termination argument: -termination_by x2' => sizeOf x2' +termination_by x2' => x2' Inferred termination argument: -termination_by x2 => sizeOf x2 +termination_by x2 => x2 +Inferred termination argument: +termination_by x1 => sizeOf x1 Inferred termination argument: termination_by x2 => SizeOf.sizeOf x2 Inferred termination argument: termination_by x1 x2 => SizeOf.sizeOf x1 Inferred termination argument: termination_by x2 => SizeOf.sizeOf x2 +Inferred termination argument: +termination_by x1 => x1 +Inferred termination argument: +termination_by x1 => sizeOf x1 +Inferred termination argument: +termination_by x1 => x1 +Inferred termination argument: +termination_by sizeOf o +guessLex.lean:217:0-229:3: error: Could not find a decreasing measure. +The arguments relate at each recursive call as follows: +(<, ≤, =: relation proved, ? all proofs failed, _: no proof attempted) +Call from MutualNotNat2.MutualNotNat2.MutualNotNat3.foo to MutualNotNat2.MutualNotNat2.MutualNotNat3.bar at 221:23-35: + x1 +x1 < +Call from MutualNotNat2.MutualNotNat2.MutualNotNat3.bar to MutualNotNat2.MutualNotNat2.MutualNotNat3.foo at 226:30-42: + x1 +x1 ? + +Please use `termination_by` to specify a decreasing measure. diff --git a/tests/lean/guessLexTricky.lean.expected.out b/tests/lean/guessLexTricky.lean.expected.out index 6e98a304c024..816a3a696acd 100644 --- a/tests/lean/guessLexTricky.lean.expected.out +++ b/tests/lean/guessLexTricky.lean.expected.out @@ -1,6 +1,6 @@ Inferred termination argument: -termination_by (sizeOf y, 1, 0) +termination_by (y, 1, 0) Inferred termination argument: -termination_by (sizeOf y, 0, 1) +termination_by (y, 0, 1) Inferred termination argument: -termination_by (sizeOf y, 0, 0) +termination_by (y, 0, 0) diff --git a/tests/lean/guessLexTricky2.lean.expected.out b/tests/lean/guessLexTricky2.lean.expected.out index 24db6dc1bfbe..d27a0e6ab4bf 100644 --- a/tests/lean/guessLexTricky2.lean.expected.out +++ b/tests/lean/guessLexTricky2.lean.expected.out @@ -1,4 +1,4 @@ Inferred termination argument: -termination_by x1 x2 x3 => (sizeOf x1, sizeOf x2, 0) +termination_by x1 x2 x3 => (x1, x2, 0) Inferred termination argument: -termination_by x1 x2 x3 => (sizeOf x1, sizeOf x2, 1) +termination_by x1 x2 x3 => (x1, x2, 1) diff --git a/tests/lean/interactive/terminationBySuggestion.lean b/tests/lean/interactive/terminationBySuggestion.lean index 928c0cc081fb..cb6480bcc0a8 100644 --- a/tests/lean/interactive/terminationBySuggestion.lean +++ b/tests/lean/interactive/terminationBySuggestion.lean @@ -1,7 +1,19 @@ - def ackermann (n m : Nat) := match n, m with | 0, m => m + 1 | .succ n, 0 => ackermann n 1 | .succ n, .succ m => ackermann n (ackermann (n + 1) m) termination_by? --^ codeAction + +-- Check hat we print this even if there is only one plausible measure +def onlyOneMeasure (n : Nat) := match n with + | 0 => 0 + | .succ n => onlyOneMeasure n +termination_by? + --^ codeAction + +def anonymousMeasure : Nat → Nat + | 0 => 0 + | .succ n => anonymousMeasure n +termination_by? + --^ codeAction diff --git a/tests/lean/interactive/terminationBySuggestion.lean.expected.out b/tests/lean/interactive/terminationBySuggestion.lean.expected.out index a1c73f09d83f..b0624016ed71 100644 --- a/tests/lean/interactive/terminationBySuggestion.lean.expected.out +++ b/tests/lean/interactive/terminationBySuggestion.lean.expected.out @@ -1,4 +1,4 @@ -{"title": "Try this: termination_by (sizeOf n, sizeOf m)", +{"title": "Try this: termination_by (n, m)", "kind": "quickfix", "isPreferred": true, "edit": @@ -7,6 +7,30 @@ {"version": 1, "uri": "file:///terminationBySuggestion.lean"}, "edits": [{"range": - {"start": {"line": 5, "character": 0}, - "end": {"line": 5, "character": 15}}, - "newText": "termination_by (sizeOf n, sizeOf m)"}]}]}} + {"start": {"line": 4, "character": 0}, + "end": {"line": 4, "character": 15}}, + "newText": "termination_by (n, m)"}]}]}} +{"title": "Try this: termination_by n", + "kind": "quickfix", + "isPreferred": true, + "edit": + {"documentChanges": + [{"textDocument": + {"version": 1, "uri": "file:///terminationBySuggestion.lean"}, + "edits": + [{"range": + {"start": {"line": 11, "character": 0}, + "end": {"line": 11, "character": 15}}, + "newText": "termination_by n"}]}]}} +{"title": "Try this: termination_by x1 => x1", + "kind": "quickfix", + "isPreferred": true, + "edit": + {"documentChanges": + [{"textDocument": + {"version": 1, "uri": "file:///terminationBySuggestion.lean"}, + "edits": + [{"range": + {"start": {"line": 17, "character": 0}, + "end": {"line": 17, "character": 15}}, + "newText": "termination_by x1 => x1"}]}]}} diff --git a/tests/lean/issue2981.lean.expected.out b/tests/lean/issue2981.lean.expected.out index f185b622f2de..6c181f508042 100644 --- a/tests/lean/issue2981.lean.expected.out +++ b/tests/lean/issue2981.lean.expected.out @@ -3,11 +3,11 @@ Tactic is run (ideally only twice) Tactic is run (ideally only twice) Tactic is run (ideally only once, in most general context) n : Nat -⊢ (invImage (fun a => sizeOf a) instWellFoundedRelation).1 n (Nat.succ n) +⊢ (invImage (fun a => a) instWellFoundedRelation).1 n (Nat.succ n) Tactic is run (ideally only twice, in most general context) Tactic is run (ideally only twice, in most general context) n : Nat ⊢ sizeOf n < sizeOf (Nat.succ n) n m : Nat -⊢ (invImage (fun a => PSigma.casesOn a fun x1 snd => sizeOf x1) instWellFoundedRelation).1 { fst := n, snd := m + 1 } +⊢ (invImage (fun a => PSigma.casesOn a fun x1 snd => x1) instWellFoundedRelation).1 { fst := n, snd := m + 1 } { fst := Nat.succ n, snd := m } From ebefee0b7ddaf42d02ee4963b4a3d90d97ea5c8b Mon Sep 17 00:00:00 2001 From: Mac Malone Date: Mon, 11 Mar 2024 12:14:24 -0400 Subject: [PATCH 11/32] chore: response file to avoid arg limits in lean static lib build (#3612) --- src/lean.mk.in | 9 +++++++++ stage0/src/lean.mk.in | 9 +++++++++ 2 files changed, 18 insertions(+) diff --git a/src/lean.mk.in b/src/lean.mk.in index d6ed3e021a83..60fc1241a66e 100644 --- a/src/lean.mk.in +++ b/src/lean.mk.in @@ -123,9 +123,18 @@ else $(LEANC) -o "$@" $^ $(LEANC_OPTS) $(LINK_OPTS) endif +ifeq (@CMAKE_SYSTEM_NAME@, Windows) $(LIB_OUT)/$(STATIC_LIB_NAME): $(addprefix $(TEMP_OUT)/,$(REL_OS)) | $(LIB_OUT) @rm -f $@ + $(file >$@.in) $(foreach O,$^,$(file >>$@.in,"$O")) + @$(LEAN_AR) rcs $@ @$@.in + @rm -f $@.in +else +$(LIB_OUT)/$(STATIC_LIB_NAME): $(addprefix $(TEMP_OUT)/,$(REL_OS)) | $(LIB_OUT) + @rm -f $@ +# no response file support on macOS, but also no need for them @$(LEAN_AR) rcs $@ $^ +endif clean: rm -rf $(OUT) diff --git a/stage0/src/lean.mk.in b/stage0/src/lean.mk.in index d6ed3e021a83..60fc1241a66e 100644 --- a/stage0/src/lean.mk.in +++ b/stage0/src/lean.mk.in @@ -123,9 +123,18 @@ else $(LEANC) -o "$@" $^ $(LEANC_OPTS) $(LINK_OPTS) endif +ifeq (@CMAKE_SYSTEM_NAME@, Windows) $(LIB_OUT)/$(STATIC_LIB_NAME): $(addprefix $(TEMP_OUT)/,$(REL_OS)) | $(LIB_OUT) @rm -f $@ + $(file >$@.in) $(foreach O,$^,$(file >>$@.in,"$O")) + @$(LEAN_AR) rcs $@ @$@.in + @rm -f $@.in +else +$(LIB_OUT)/$(STATIC_LIB_NAME): $(addprefix $(TEMP_OUT)/,$(REL_OS)) | $(LIB_OUT) + @rm -f $@ +# no response file support on macOS, but also no need for them @$(LEAN_AR) rcs $@ $^ +endif clean: rm -rf $(OUT) From d9b6794e2fa684a72b3ae369b05bec71305521b3 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Mon, 11 Mar 2024 17:29:56 +0100 Subject: [PATCH 12/32] refactor: termination_by parser to use binderIdent (#3652) this way we should be able to use `elabBinders` to parse the binders. --- src/Lean/Parser/Term.lean | 102 ++++++++++++++++++++------------------ stage0/src/stdlib_flags.h | 2 +- 2 files changed, 54 insertions(+), 50 deletions(-) diff --git a/src/Lean/Parser/Term.lean b/src/Lean/Parser/Term.lean index 0fc9514dc129..a4b2b2ff8c01 100644 --- a/src/Lean/Parser/Term.lean +++ b/src/Lean/Parser/Term.lean @@ -97,55 +97,6 @@ end Tactic def darrow : Parser := " => " def semicolonOrLinebreak := ";" <|> checkLinebreakBefore >> pushNone -namespace Termination - -/- -Termination suffix parsers, typically thought of as part of a command, but due to -letrec we need them here already. --/ - -/-- -Specify a termination argument for well-founded termination: -``` -termination_by a - b -``` -indicates that termination of the currently defined recursive function follows -because the difference between the the arguments `a` and `b`. - -If the fuction takes further argument after the colon, you can name them as follows: -``` -def example (a : Nat) : Nat → Nat → Nat := -termination_by b c => a - b -``` - -If omitted, a termination argument will be inferred. If written as `termination_by?`, -the inferrred termination argument will be suggested. --/ -def terminationBy := leading_parser - "termination_by " >> - optional (atomic (many (ppSpace >> (ident <|> "_")) >> " => ")) >> - termParser - -@[inherit_doc terminationBy] -def terminationBy? := leading_parser - "termination_by?" - -/-- -Manually prove that the termination argument (as specified with `termination_by` or inferred) -decreases at each recursive call. - -By default, the tactic `decreasing_tactic` is used. --/ -def decreasingBy := leading_parser - ppDedent ppLine >> "decreasing_by " >> Tactic.tacticSeqIndentGt - -/-- -Termination hints are `termination_by` and `decreasing_by`, in that order. --/ -def suffix := leading_parser - optional (ppDedent ppLine >> (terminationBy? <|> terminationBy)) >> optional decreasingBy - -end Termination namespace Term @@ -596,6 +547,59 @@ def attrInstance := ppGroup $ leading_parser attrKind >> attrParser def attributes := leading_parser "@[" >> withoutPosition (sepBy1 attrInstance ", ") >> "] " + +end Term +namespace Termination + +/- +Termination suffix parsers, typically thought of as part of a command, but due to +letrec we need them here already. +-/ + +/-- +Specify a termination argument for well-founded termination: +``` +termination_by a - b +``` +indicates that termination of the currently defined recursive function follows +because the difference between the the arguments `a` and `b`. + +If the fuction takes further argument after the colon, you can name them as follows: +``` +def example (a : Nat) : Nat → Nat → Nat := +termination_by b c => a - b +``` + +If omitted, a termination argument will be inferred. If written as `termination_by?`, +the inferrred termination argument will be suggested. +-/ +def terminationBy := leading_parser + "termination_by " >> + optional (atomic (many (ppSpace >> Term.binderIdent) >> " => ")) >> + termParser + +@[inherit_doc terminationBy] +def terminationBy? := leading_parser + "termination_by?" + +/-- +Manually prove that the termination argument (as specified with `termination_by` or inferred) +decreases at each recursive call. + +By default, the tactic `decreasing_tactic` is used. +-/ +def decreasingBy := leading_parser + ppDedent ppLine >> "decreasing_by " >> Tactic.tacticSeqIndentGt + +/-- +Termination hints are `termination_by` and `decreasing_by`, in that order. +-/ +def suffix := leading_parser + optional (ppDedent ppLine >> (terminationBy? <|> terminationBy)) >> optional decreasingBy + +end Termination +namespace Term + /-- `letRecDecl` matches the body of a let-rec declaration: a doc comment, attributes, and then a let declaration without the `let` keyword, such as `/-- foo -/ @[simp] bar := 1`. -/ def letRecDecl := leading_parser diff --git a/stage0/src/stdlib_flags.h b/stage0/src/stdlib_flags.h index 0699845ba452..658ab0874e68 100644 --- a/stage0/src/stdlib_flags.h +++ b/stage0/src/stdlib_flags.h @@ -8,7 +8,7 @@ options get_default_options() { // switch to `true` for ABI-breaking changes affecting meta code opts = opts.update({"interpreter", "prefer_native"}, false); // switch to `true` for changing built-in parsers used in quotations - opts = opts.update({"internal", "parseQuotWithCurrentStage"}, false); + opts = opts.update({"internal", "parseQuotWithCurrentStage"}, true); // toggling `parseQuotWithCurrentStage` may also require toggling the following option if macros/syntax // with custom precheck hooks were affected opts = opts.update({"quotPrecheck"}, true); From 1388f6bc83b5ea085fe14faa1db213e745c3e398 Mon Sep 17 00:00:00 2001 From: Lean stage0 autoupdater <> Date: Mon, 11 Mar 2024 17:22:37 +0000 Subject: [PATCH 13/32] chore: update stage0 --- stage0/src/runtime/compact.cpp | 2 +- stage0/src/runtime/thread.cpp | 14 +- stage0/src/runtime/thread.h | 5 + stage0/src/shell/CMakeLists.txt | 42 +- stage0/src/stdlib_flags.h | 2 +- stage0/stdlib/Init/Data/Nat.c | 10 +- stage0/stdlib/Init/Data/Nat/Compare.c | 33 + stage0/stdlib/Init/Data/Nat/Gcd.c | 52 +- stage0/stdlib/Init/Data/Nat/Lcm.c | 60 + stage0/stdlib/Init/Data/Nat/Lemmas.c | 6 +- stage0/stdlib/Lean/Compiler/LCNF/Main.c | 698 +- stage0/stdlib/Lean/Compiler/Main.c | 4 +- stage0/stdlib/Lean/CoreM.c | 466 +- stage0/stdlib/Lean/Elab/Command.c | 1080 +- stage0/stdlib/Lean/Elab/Deriving/DecEq.c | 4 +- stage0/stdlib/Lean/Elab/LetRec.c | 996 +- stage0/stdlib/Lean/Elab/MutualDef.c | 354 +- stage0/stdlib/Lean/Elab/PreDefinition/Main.c | 4 +- .../Lean/Elab/PreDefinition/WF/GuessLex.c | 16580 +++++++++------- .../stdlib/Lean/Elab/PreDefinition/WF/Main.c | 4 +- .../Elab/PreDefinition/WF/TerminationHint.c | 714 +- stage0/stdlib/Lean/Elab/Syntax.c | 4 +- stage0/stdlib/Lean/Elab/Tactic/Basic.c | 500 +- stage0/stdlib/Lean/Elab/Term.c | 1063 +- stage0/stdlib/Lean/Meta/Basic.c | 450 +- stage0/stdlib/Lean/Meta/ExprDefEq.c | 20 +- stage0/stdlib/Lean/Meta/Injective.c | 496 +- stage0/stdlib/Lean/Meta/SynthInstance.c | 1112 +- stage0/stdlib/Lean/Meta/Tactic/Simp/Rewrite.c | 453 +- stage0/stdlib/Lean/Parser/Term.c | 2838 ++- stage0/stdlib/Lean/PrettyPrinter.c | 11 +- .../Lean/PrettyPrinter/Delaborator/Builtins.c | 6107 ++---- stage0/stdlib/Lean/Util/Trace.c | 1392 +- .../Lean/Widget/InteractiveDiagnostic.c | 78 +- 34 files changed, 18652 insertions(+), 17002 deletions(-) create mode 100644 stage0/stdlib/Init/Data/Nat/Compare.c create mode 100644 stage0/stdlib/Init/Data/Nat/Lcm.c diff --git a/stage0/src/runtime/compact.cpp b/stage0/src/runtime/compact.cpp index b2269a4cd631..2dcd60d1e355 100644 --- a/stage0/src/runtime/compact.cpp +++ b/stage0/src/runtime/compact.cpp @@ -342,7 +342,7 @@ void object_compactor::operator()(object * o) { g_tag_counters[lean_ptr_tag(curr)]++; #endif switch (lean_ptr_tag(curr)) { - case LeanClosure: lean_internal_panic("closures cannot be compacted"); + case LeanClosure: lean_internal_panic("closures cannot be compacted. One possible cause of this error is trying to store a function in a persistent environment extension."); case LeanArray: r = insert_array(curr); break; case LeanScalarArray: insert_sarray(curr); break; case LeanString: insert_string(curr); break; diff --git a/stage0/src/runtime/thread.cpp b/stage0/src/runtime/thread.cpp index e742fe1e4184..6aab3fbc6d17 100644 --- a/stage0/src/runtime/thread.cpp +++ b/stage0/src/runtime/thread.cpp @@ -46,18 +46,26 @@ void reset_thread_local() { using runnable = std::function; -static void thread_main(void * p) { +extern "C" LEAN_EXPORT void lean_initialize_thread() { #ifdef LEAN_SMALL_ALLOCATOR init_thread_heap(); #endif +} + +extern "C" LEAN_EXPORT void lean_finalize_thread() { + run_thread_finalizers(); + run_post_thread_finalizers(); +} + +static void thread_main(void * p) { + lean_initialize_thread(); std::unique_ptr f; f.reset(reinterpret_cast(p)); (*f)(); f.reset(); - run_thread_finalizers(); - run_post_thread_finalizers(); + lean_finalize_thread(); } #if defined(LEAN_MULTI_THREAD) diff --git a/stage0/src/runtime/thread.h b/stage0/src/runtime/thread.h index decebaeae0e8..0d459282025c 100644 --- a/stage0/src/runtime/thread.h +++ b/stage0/src/runtime/thread.h @@ -217,9 +217,14 @@ static T & GETTER_NAME() { \ } namespace lean { +// module initializer pair (NOT for initializing individual threads!) void initialize_thread(); void finalize_thread(); +// thread initializer pair, for reverse FFI +extern "C" LEAN_EXPORT void lean_initialize_thread(); +extern "C" LEAN_EXPORT void lean_finalize_thread(); + typedef void (*thread_finalizer)(void *); // NOLINT LEAN_EXPORT void register_post_thread_finalizer(thread_finalizer fn, void * p); LEAN_EXPORT void register_thread_finalizer(thread_finalizer fn, void * p); diff --git a/stage0/src/shell/CMakeLists.txt b/stage0/src/shell/CMakeLists.txt index 87da1764b43b..9bcc910ad121 100644 --- a/stage0/src/shell/CMakeLists.txt +++ b/stage0/src/shell/CMakeLists.txt @@ -119,14 +119,18 @@ ENDFOREACH(T) # LEAN BENCHMARK TESTS # do not test all .lean files in bench/ -file(GLOB LEANBENCHTESTS "${LEAN_SOURCE_DIR}/../tests/bench/*.lean.expected.out") -FOREACH(T_OUT ${LEANBENCHTESTS}) - string(REPLACE ".expected.out" "" T ${T_OUT}) - GET_FILENAME_COMPONENT(T_NAME ${T} NAME) - add_test(NAME "leanbenchtest_${T_NAME}" - WORKING_DIRECTORY "${LEAN_SOURCE_DIR}/../tests/bench" - COMMAND bash -c "${TEST_VARS} ./test_single.sh ${T_NAME}") -ENDFOREACH(T_OUT) +if(${CMAKE_SYSTEM_NAME} MATCHES "Windows") + message(STATUS "Skipping compiler tests on Windows because of shared library limit on number of exported symbols") +else() + file(GLOB LEANBENCHTESTS "${LEAN_SOURCE_DIR}/../tests/bench/*.lean.expected.out") + FOREACH(T_OUT ${LEANBENCHTESTS}) + string(REPLACE ".expected.out" "" T ${T_OUT}) + GET_FILENAME_COMPONENT(T_NAME ${T} NAME) + add_test(NAME "leanbenchtest_${T_NAME}" + WORKING_DIRECTORY "${LEAN_SOURCE_DIR}/../tests/bench" + COMMAND bash -c "${TEST_VARS} ./test_single.sh ${T_NAME}") + ENDFOREACH(T_OUT) +endif() file(GLOB LEANINTERPTESTS "${LEAN_SOURCE_DIR}/../tests/plugin/*.lean") FOREACH(T ${LEANINTERPTESTS}) @@ -146,15 +150,19 @@ FOREACH(T ${LEANT0TESTS}) ENDFOREACH(T) # LEAN PACKAGE TESTS -file(GLOB LEANPKGTESTS "${LEAN_SOURCE_DIR}/../tests/pkg/*") -FOREACH(T ${LEANPKGTESTS}) - if(IS_DIRECTORY ${T}) - GET_FILENAME_COMPONENT(T_NAME ${T} NAME) - add_test(NAME "leanpkgtest_${T_NAME}" - WORKING_DIRECTORY "${T}" - COMMAND bash -c "${TEST_VARS} ./test.sh") - endif() -ENDFOREACH(T) +if(${CMAKE_SYSTEM_NAME} MATCHES "Windows") + message(STATUS "Skipping compiler tests on Windows because of shared library limit on number of exported symbols") +else() + file(GLOB LEANPKGTESTS "${LEAN_SOURCE_DIR}/../tests/pkg/*") + FOREACH(T ${LEANPKGTESTS}) + if(IS_DIRECTORY ${T}) + GET_FILENAME_COMPONENT(T_NAME ${T} NAME) + add_test(NAME "leanpkgtest_${T_NAME}" + WORKING_DIRECTORY "${T}" + COMMAND bash -c "${TEST_VARS} ./test.sh") + endif() + ENDFOREACH(T) +endif() # LEAN SERVER TESTS file(GLOB LEANTESTS "${LEAN_SOURCE_DIR}/../tests/lean/server/*.lean") diff --git a/stage0/src/stdlib_flags.h b/stage0/src/stdlib_flags.h index 658ab0874e68..0699845ba452 100644 --- a/stage0/src/stdlib_flags.h +++ b/stage0/src/stdlib_flags.h @@ -8,7 +8,7 @@ options get_default_options() { // switch to `true` for ABI-breaking changes affecting meta code opts = opts.update({"interpreter", "prefer_native"}, false); // switch to `true` for changing built-in parsers used in quotations - opts = opts.update({"internal", "parseQuotWithCurrentStage"}, true); + opts = opts.update({"internal", "parseQuotWithCurrentStage"}, false); // toggling `parseQuotWithCurrentStage` may also require toggling the following option if macros/syntax // with custom precheck hooks were affected opts = opts.update({"quotPrecheck"}, true); diff --git a/stage0/stdlib/Init/Data/Nat.c b/stage0/stdlib/Init/Data/Nat.c index 03648fc804aa..fcd9a417cc00 100644 --- a/stage0/stdlib/Init/Data/Nat.c +++ b/stage0/stdlib/Init/Data/Nat.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Init.Data.Nat -// Imports: Init.Data.Nat.Basic Init.Data.Nat.Div Init.Data.Nat.Dvd Init.Data.Nat.Gcd Init.Data.Nat.MinMax Init.Data.Nat.Bitwise Init.Data.Nat.Control Init.Data.Nat.Log2 Init.Data.Nat.Power2 Init.Data.Nat.Linear Init.Data.Nat.SOM Init.Data.Nat.Lemmas Init.Data.Nat.Mod +// Imports: Init.Data.Nat.Basic Init.Data.Nat.Div Init.Data.Nat.Dvd Init.Data.Nat.Gcd Init.Data.Nat.MinMax Init.Data.Nat.Bitwise Init.Data.Nat.Control Init.Data.Nat.Log2 Init.Data.Nat.Power2 Init.Data.Nat.Linear Init.Data.Nat.SOM Init.Data.Nat.Lemmas Init.Data.Nat.Mod Init.Data.Nat.Lcm Init.Data.Nat.Compare #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -26,6 +26,8 @@ lean_object* initialize_Init_Data_Nat_Linear(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_Nat_SOM(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_Nat_Lemmas(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_Nat_Mod(uint8_t builtin, lean_object*); +lean_object* initialize_Init_Data_Nat_Lcm(uint8_t builtin, lean_object*); +lean_object* initialize_Init_Data_Nat_Compare(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Init_Data_Nat(uint8_t builtin, lean_object* w) { lean_object * res; @@ -70,6 +72,12 @@ lean_dec_ref(res); res = initialize_Init_Data_Nat_Mod(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Init_Data_Nat_Lcm(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Init_Data_Nat_Compare(builtin, 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)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Init/Data/Nat/Compare.c b/stage0/stdlib/Init/Data/Nat/Compare.c new file mode 100644 index 000000000000..18b31703b43a --- /dev/null +++ b/stage0/stdlib/Init/Data/Nat/Compare.c @@ -0,0 +1,33 @@ +// Lean compiler output +// Module: Init.Data.Nat.Compare +// Imports: Init.Classical Init.Data.Ord +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +lean_object* initialize_Init_Classical(uint8_t builtin, lean_object*); +lean_object* initialize_Init_Data_Ord(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Init_Data_Nat_Compare(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Init_Classical(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Init_Data_Ord(builtin, 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)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Init/Data/Nat/Gcd.c b/stage0/stdlib/Init/Data/Nat/Gcd.c index 0e64f1f5320c..ff3b33818968 100644 --- a/stage0/stdlib/Init/Data/Nat/Gcd.c +++ b/stage0/stdlib/Init/Data/Nat/Gcd.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Init.Data.Nat.Gcd -// Imports: Init.Data.Nat.Dvd +// Imports: Init.Data.Nat.Dvd Init.NotationExtra Init.RCases #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -14,7 +14,11 @@ extern "C" { #endif lean_object* lean_nat_gcd(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Nat_prod__dvd__and__dvd__of__dvd__prod(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_gcd___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Nat_prod__dvd__and__dvd__of__dvd__prod___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_nat_div(lean_object*, lean_object*); +uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_gcd___boxed(lean_object* x_1, lean_object* x_2) { _start: { @@ -25,7 +29,47 @@ lean_dec(x_1); return x_3; } } +LEAN_EXPORT lean_object* l_Nat_prod__dvd__and__dvd__of__dvd__prod(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_nat_gcd(x_1, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_nat_dec_eq(x_5, x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; +lean_dec(x_3); +x_8 = lean_nat_div(x_1, x_5); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_5); +lean_ctor_set(x_9, 1, x_8); +return x_9; +} +else +{ +lean_object* x_10; +lean_dec(x_5); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_6); +lean_ctor_set(x_10, 1, x_3); +return x_10; +} +} +} +LEAN_EXPORT lean_object* l_Nat_prod__dvd__and__dvd__of__dvd__prod___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_Nat_prod__dvd__and__dvd__of__dvd__prod(x_1, x_2, x_3, x_4); +lean_dec(x_2); +lean_dec(x_1); +return x_5; +} +} lean_object* initialize_Init_Data_Nat_Dvd(uint8_t builtin, lean_object*); +lean_object* initialize_Init_NotationExtra(uint8_t builtin, lean_object*); +lean_object* initialize_Init_RCases(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Init_Data_Nat_Gcd(uint8_t builtin, lean_object* w) { lean_object * res; @@ -34,6 +78,12 @@ _G_initialized = true; res = initialize_Init_Data_Nat_Dvd(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Init_NotationExtra(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Init_RCases(builtin, 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)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Init/Data/Nat/Lcm.c b/stage0/stdlib/Init/Data/Nat/Lcm.c new file mode 100644 index 000000000000..53899834d429 --- /dev/null +++ b/stage0/stdlib/Init/Data/Nat/Lcm.c @@ -0,0 +1,60 @@ +// Lean compiler output +// Module: Init.Data.Nat.Lcm +// Imports: Init.Data.Nat.Gcd Init.Data.Nat.Lemmas +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +lean_object* lean_nat_gcd(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Nat_lcm(lean_object*, lean_object*); +lean_object* lean_nat_div(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Nat_lcm___boxed(lean_object*, lean_object*); +lean_object* lean_nat_mul(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Nat_lcm(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = lean_nat_mul(x_1, x_2); +x_4 = lean_nat_gcd(x_1, x_2); +x_5 = lean_nat_div(x_3, x_4); +lean_dec(x_4); +lean_dec(x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Nat_lcm___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Nat_lcm(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +lean_object* initialize_Init_Data_Nat_Gcd(uint8_t builtin, lean_object*); +lean_object* initialize_Init_Data_Nat_Lemmas(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Init_Data_Nat_Lcm(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Init_Data_Nat_Gcd(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Init_Data_Nat_Lemmas(builtin, 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)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Init/Data/Nat/Lemmas.c b/stage0/stdlib/Init/Data/Nat/Lemmas.c index a3079f968d49..5c16deb7f796 100644 --- a/stage0/stdlib/Init/Data/Nat/Lemmas.c +++ b/stage0/stdlib/Init/Data/Nat/Lemmas.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Init.Data.Nat.Lemmas -// Imports: Init.Data.Nat.Dvd Init.Data.Nat.MinMax Init.Data.Nat.Log2 Init.Data.Nat.Power2 +// Imports: Init.Data.Nat.Dvd Init.Data.Nat.MinMax Init.Data.Nat.Log2 Init.Data.Nat.Power2 Init.Omega #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -273,6 +273,7 @@ lean_object* initialize_Init_Data_Nat_Dvd(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_Nat_MinMax(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_Nat_Log2(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_Nat_Power2(uint8_t builtin, lean_object*); +lean_object* initialize_Init_Omega(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Init_Data_Nat_Lemmas(uint8_t builtin, lean_object* w) { lean_object * res; @@ -290,6 +291,9 @@ lean_dec_ref(res); res = initialize_Init_Data_Nat_Power2(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Init_Omega(builtin, 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)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Compiler/LCNF/Main.c b/stage0/stdlib/Lean/Compiler/LCNF/Main.c index 9df0b4cd1d22..50cfdcc52d03 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/Main.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/Main.c @@ -13,14 +13,14 @@ #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*); static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1408____closed__21; static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__19; static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1408____closed__2; uint8_t lean_is_matcher(lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1408_(lean_object*); -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_PassManager_run___spec__7___lambda__1___closed__11; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__4; +LEAN_EXPORT lean_object* l_Lean_addTrace___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* l_Lean_Compiler_LCNF_toDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_main___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_profiler; @@ -29,19 +29,19 @@ lean_object* l_Lean_Compiler_LCNF_PassManager_run___lambda__3___boxed(lean_objec lean_object* lean_mk_empty_array_with_capacity(lean_object*); 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_PassManager_run___spec__7___lambda__1___closed__1; 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_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; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__4(lean_object*, lean_object*, uint8_t, lean_object*, 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_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*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__9___closed__2; 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_1408____closed__6; lean_object* l_Lean_MessageData_ofList(lean_object*); @@ -50,88 +50,80 @@ lean_object* l_Lean_PersistentArray_push___rarg(lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); 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_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_object* l_Lean_Compiler_LCNF_PassManager_run___lambda__3(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_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*); static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1408____closed__8; lean_object* l_Lean_Compiler_LCNF_PassManager_run(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1408____closed__20; -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__7; lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_float_decLt(double, double); 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_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__5___closed__1; 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*); static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1408____closed__1; 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_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; -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__16; 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*); 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*); static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__15; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__5(lean_object*, 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_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; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__10; 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_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__9___closed__1; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__30; static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1408____closed__12; lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___lambda__3(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_1408____closed__23; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___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* 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_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_object*); 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_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__4; 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_1408____closed__7; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__13; static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__23; 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__7___lambda__1___closed__19; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__9; 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*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); lean_object* l_Lean_Name_append(lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__16; 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; @@ -139,11 +131,12 @@ static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main__ 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_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__12; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__15; 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*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__9; 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; @@ -151,48 +144,53 @@ lean_object* l_Lean_Compiler_LCNF_showDecl___boxed(lean_object*, lean_object*, l 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*); lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__3; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1(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*, uint8_t, lean_object*, lean_object*, uint8_t, double, lean_object*, 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_shouldGenerateCode___closed__9; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__7; +LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_LCNF_PassManager_run___spec__8(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__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_EXPORT lean_object* l_Array_forInUnsafe_loop___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_object*, lean_object*); lean_object* l_Lean_KVMap_setBool(lean_object*, lean_object*, uint8_t); +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Compiler_LCNF_PassManager_run___spec__6___boxed(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__9___closed__3; 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_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__9___closed__4; +lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNode___spec__1(size_t, size_t, lean_object*); static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__18; -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; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__14; lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode_isCompIrrelevant(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_1408____closed__9; -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__10___closed__1; static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1408____closed__3; -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__15; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__2; static lean_object* l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__4___closed__2; lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -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; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__11(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_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_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_object* l_Lean_Compiler_LCNF_PassManager_run___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_PassManager_run___spec__10(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__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__8(lean_object*, size_t, size_t, 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__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_float_to_string(double); 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*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__8; 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*); static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1408____closed__22; 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_PassManager_run___spec__7___lambda__1___closed__18; 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*); @@ -200,7 +198,6 @@ static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint_ 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; -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_shouldGenerateCode___closed__14; @@ -209,10 +206,9 @@ lean_object* lean_nat_sub(lean_object*, lean_object*); 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_1408____closed__15; -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_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__17; 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*, lean_object*, lean_object*); @@ -227,33 +223,30 @@ 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_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; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__12; lean_object* l_Lean_Compiler_LCNF_ppDecl_x27(lean_object*, lean_object*, lean_object*, lean_object*); 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); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___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*); 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_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_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1408____closed__4; static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__25; 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; static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1408____closed__16; -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*); static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1408____closed__11; 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; static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1408____closed__17; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___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); +lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__13___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__12; static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1408____closed__19; lean_object* l_Lean_Compiler_LCNF_getDeclAt_x3f(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); @@ -264,22 +257,26 @@ static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main__ static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1408____closed__13; lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__4; static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1408____closed__5; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7(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__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*); static lean_object* l_Lean_Compiler_LCNF_PassManager_run___lambda__2___closed__2; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__6; static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__5; lean_object* l_Nat_repr(lean_object*); lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1408____closed__14; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__5; lean_object* l_List_toArrayAux___rarg(lean_object*, lean_object*); -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_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___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_Lean_Meta_isMatcher___at_Lean_Compiler_LCNF_shouldGenerateCode___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___boxed(lean_object*, lean_object*, lean_object*, 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; lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode_isCompIrrelevant(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { @@ -2578,249 +2575,153 @@ 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) { +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: { -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_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; 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; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_11 = lean_ctor_get(x_8, 2); +x_12 = lean_st_ref_get(x_9, x_10); +x_13 = lean_ctor_get(x_12, 0); 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_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = lean_ctor_get(x_13, 3); +lean_inc(x_15); +lean_dec(x_13); 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_20 = l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNode___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); +x_22 = lean_st_ref_get(x_9, x_14); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = lean_st_ref_get(x_7, x_24); +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); lean_dec(x_26); -x_27 = lean_box(0); -lean_ctor_set(x_24, 0, x_27); -return x_24; +x_29 = lean_ctor_get(x_27, 0); +lean_inc(x_29); +lean_dec(x_27); +x_30 = l_Lean_Compiler_LCNF_LCtx_toLocalContext(x_29); +x_31 = l_Lean_Compiler_LCNF_shouldGenerateCode___closed__14; +lean_inc(x_11); +x_32 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_32, 0, x_25); +lean_ctor_set(x_32, 1, x_31); +lean_ctor_set(x_32, 2, x_30); +lean_ctor_set(x_32, 3, x_11); +x_33 = lean_alloc_ctor(3, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_21); +x_34 = lean_st_ref_take(x_9, x_28); +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_34, 1); +lean_inc(x_36); +lean_dec(x_34); +x_37 = !lean_is_exclusive(x_35); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; +x_38 = lean_ctor_get(x_35, 3); +lean_dec(x_38); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_3); +lean_ctor_set(x_39, 1, x_33); +x_40 = l_Lean_PersistentArray_push___rarg(x_1, x_39); +lean_ctor_set(x_35, 3, x_40); +x_41 = lean_st_ref_set(x_9, x_35, x_36); +x_42 = !lean_is_exclusive(x_41); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; +x_43 = lean_ctor_get(x_41, 0); +lean_dec(x_43); +x_44 = lean_box(0); +lean_ctor_set(x_41, 0, x_44); +return x_41; } 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_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_41, 1); +lean_inc(x_45); +lean_dec(x_41); +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_31; lean_object* x_32; lean_object* x_33; 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_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_48 = lean_ctor_get(x_35, 0); +x_49 = lean_ctor_get(x_35, 1); +x_50 = lean_ctor_get(x_35, 2); +x_51 = lean_ctor_get(x_35, 4); +x_52 = lean_ctor_get(x_35, 5); +x_53 = lean_ctor_get(x_35, 6); +lean_inc(x_53); +lean_inc(x_52); +lean_inc(x_51); +lean_inc(x_50); +lean_inc(x_49); 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; +lean_dec(x_35); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_3); +lean_ctor_set(x_54, 1, x_33); +x_55 = l_Lean_PersistentArray_push___rarg(x_1, x_54); +x_56 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_56, 0, x_48); +lean_ctor_set(x_56, 1, x_49); +lean_ctor_set(x_56, 2, x_50); +lean_ctor_set(x_56, 3, x_55); +lean_ctor_set(x_56, 4, x_51); +lean_ctor_set(x_56, 5, x_52); +lean_ctor_set(x_56, 6, x_53); +x_57 = lean_st_ref_set(x_9, x_56, x_36); +x_58 = lean_ctor_get(x_57, 1); +lean_inc(x_58); +if (lean_is_exclusive(x_57)) { + lean_ctor_release(x_57, 0); + lean_ctor_release(x_57, 1); + x_59 = x_57; } else { - lean_dec_ref(x_47); - x_49 = lean_box(0); + lean_dec_ref(x_57); + x_59 = lean_box(0); } -x_50 = lean_box(0); -if (lean_is_scalar(x_49)) { - x_51 = lean_alloc_ctor(0, 2, 0); +x_60 = lean_box(0); +if (lean_is_scalar(x_59)) { + x_61 = 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; -} + x_61 = x_59; } -} -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_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; 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; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_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); -x_39 = lean_ctor_get_uint8(x_8, sizeof(void*)*11); -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_40 = l_Lean_replaceRef(x_3, x_33); -lean_dec(x_33); -lean_inc(x_30); -x_41 = lean_alloc_ctor(0, 11, 1); -lean_ctor_set(x_41, 0, x_28); -lean_ctor_set(x_41, 1, x_29); -lean_ctor_set(x_41, 2, x_30); -lean_ctor_set(x_41, 3, x_31); -lean_ctor_set(x_41, 4, x_32); -lean_ctor_set(x_41, 5, x_40); -lean_ctor_set(x_41, 6, x_34); -lean_ctor_set(x_41, 7, x_35); -lean_ctor_set(x_41, 8, x_36); -lean_ctor_set(x_41, 9, x_37); -lean_ctor_set(x_41, 10, x_38); -lean_ctor_set_uint8(x_41, sizeof(void*)*11, x_39); -x_42 = lean_st_ref_get(x_9, 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_ctor_get(x_43, 0); -lean_inc(x_45); -lean_dec(x_43); -x_46 = lean_st_ref_get(x_7, x_44); -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_ctor_get(x_47, 0); -lean_inc(x_49); -lean_dec(x_47); -x_50 = l_Lean_Compiler_LCNF_LCtx_toLocalContext(x_49); -x_51 = l_Lean_Compiler_LCNF_shouldGenerateCode___closed__14; -x_52 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_52, 0, x_45); -lean_ctor_set(x_52, 1, x_51); -lean_ctor_set(x_52, 2, x_50); -lean_ctor_set(x_52, 3, x_30); -x_53 = lean_alloc_ctor(3, 2, 0); -lean_ctor_set(x_53, 0, x_52); -lean_ctor_set(x_53, 1, x_4); -x_54 = l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Compiler_LCNF_PassManager_run___spec__6(x_1, x_2, x_3, x_53, x_5, x_6, x_7, x_41, x_9, x_48); -lean_dec(x_41); -return x_54; +lean_ctor_set(x_61, 0, x_60); +lean_ctor_set(x_61, 1, x_58); +return x_61; } } } -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) { +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Compiler_LCNF_PassManager_run___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) @@ -2915,13 +2816,11 @@ LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManage _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); +x_15 = l_MonadExcept_ofExcept___at_Lean_Compiler_LCNF_PassManager_run___spec__6(x_5, x_8, x_9, x_10, x_11, x_14); return x_15; } } @@ -3083,6 +2982,7 @@ lean_inc(x_19); lean_dec(x_17); x_20 = l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3(x_1, x_2, x_16, x_3, x_4, x_5, x_6, x_7, lean_box(0), x_18, x_11, x_12, x_13, x_14, x_19); lean_dec(x_14); +lean_dec(x_13); lean_dec(x_11); lean_dec(x_4); return x_20; @@ -3096,6 +2996,7 @@ lean_dec(x_17); x_22 = l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__4___closed__2; x_23 = l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3(x_1, x_2, x_16, x_3, x_4, x_5, x_6, x_7, lean_box(0), x_22, x_11, x_12, x_13, x_14, x_21); lean_dec(x_14); +lean_dec(x_13); lean_dec(x_11); lean_dec(x_4); return x_23; @@ -3187,7 +3088,7 @@ x_30 = lean_st_ref_set(x_12, x_25, x_26); x_31 = lean_ctor_get(x_30, 1); lean_inc(x_31); lean_dec(x_30); -x_32 = l_MonadExcept_ofExcept___at_Lean_Compiler_LCNF_PassManager_run___spec__7(x_21, x_9, x_10, x_11, x_12, x_31); +x_32 = l_MonadExcept_ofExcept___at_Lean_Compiler_LCNF_PassManager_run___spec__6(x_21, x_9, x_10, x_11, x_12, x_31); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -3269,7 +3170,7 @@ x_50 = lean_st_ref_set(x_12, x_49, x_26); x_51 = lean_ctor_get(x_50, 1); lean_inc(x_51); lean_dec(x_50); -x_52 = l_MonadExcept_ofExcept___at_Lean_Compiler_LCNF_PassManager_run___spec__7(x_21, x_9, x_10, x_11, x_12, x_51); +x_52 = l_MonadExcept_ofExcept___at_Lean_Compiler_LCNF_PassManager_run___spec__6(x_21, x_9, x_10, x_11, x_12, x_51); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -3474,7 +3375,7 @@ return x_32; } } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__1() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__1() { _start: { lean_object* x_1; @@ -3482,16 +3383,16 @@ 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() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___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_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___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() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__3() { _start: { lean_object* x_1; @@ -3499,16 +3400,16 @@ 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() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___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_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___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() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__5() { _start: { lean_object* x_1; @@ -3516,51 +3417,51 @@ 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() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___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_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___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() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___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_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___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() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___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_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__2; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___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() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___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_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__8; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___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() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__10() { _start: { lean_object* x_1; @@ -3568,51 +3469,51 @@ 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() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___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_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___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() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___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_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___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() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___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_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__2; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___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() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___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_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__13; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___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() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__15() { _start: { lean_object* x_1; @@ -3620,51 +3521,51 @@ 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__8___lambda__1___closed__16() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___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_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___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() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___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_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___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() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___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_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__2; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___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() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___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_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__18; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___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) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___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; @@ -3677,7 +3578,7 @@ 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_11 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___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); @@ -3693,7 +3594,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_16 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___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); @@ -3709,7 +3610,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_21 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___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); @@ -3725,7 +3626,7 @@ return x_25; } } } -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) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { uint8_t x_10; @@ -3747,7 +3648,7 @@ else 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); 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); +x_13 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___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); @@ -3857,7 +3758,7 @@ return x_40; } } } -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) { +LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_LCNF_PassManager_run___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) { _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; @@ -4001,7 +3902,7 @@ return x_56; } } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__10___closed__1() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__9___closed__1() { _start: { lean_object* x_1; @@ -4009,7 +3910,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__10___closed__2() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__9___closed__2() { _start: { lean_object* x_1; @@ -4017,7 +3918,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__10___closed__3() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__9___closed__3() { _start: { lean_object* x_1; @@ -4025,20 +3926,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__10___closed__4() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__9___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__10___closed__1; -x_2 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__10___closed__2; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__9___closed__1; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__9___closed__2; x_3 = lean_unsigned_to_nat(82u); x_4 = lean_unsigned_to_nat(52u); -x_5 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__10___closed__3; +x_5 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__9___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__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) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__9(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; @@ -4074,7 +3975,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__10___closed__4; +x_26 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__9___closed__4; lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); @@ -4163,7 +4064,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__9(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__8(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); @@ -4217,7 +4118,7 @@ goto _start; } } } -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) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_PassManager_run___spec__10(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; @@ -4304,7 +4205,7 @@ return x_28; } } } -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) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__11(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; @@ -4340,7 +4241,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__10___closed__4; +x_26 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__9___closed__4; lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); @@ -4429,7 +4330,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__9(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__8(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); @@ -4546,7 +4447,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__8(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__7(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) { @@ -4600,7 +4501,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__10(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__9(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; @@ -4740,7 +4641,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__8(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__7(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) { @@ -4794,7 +4695,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__12(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__11(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; @@ -4951,7 +4852,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__11(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__10(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) { @@ -5056,7 +4957,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__11(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__10(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) { @@ -5205,7 +5106,7 @@ lean_dec(x_7); x_84 = l_Lean_Compiler_LCNF_shouldGenerateCode___closed__9; lean_inc(x_5); lean_inc(x_70); -x_85 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_PassManager_run___spec__11(x_1, x_82, x_83, x_84, x_2, x_3, x_70, x_5, x_6); +x_85 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_PassManager_run___spec__10(x_1, x_82, x_83, x_84, x_2, x_3, x_70, x_5, x_6); lean_dec(x_1); if (lean_obj_tag(x_85) == 0) { @@ -5325,7 +5226,7 @@ lean_dec(x_7); x_106 = l_Lean_Compiler_LCNF_shouldGenerateCode___closed__9; lean_inc(x_5); lean_inc(x_92); -x_107 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_PassManager_run___spec__11(x_1, x_104, x_105, x_106, x_2, x_3, x_92, x_5, x_6); +x_107 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_PassManager_run___spec__10(x_1, x_104, x_105, x_106, x_2, x_3, x_92, x_5, x_6); lean_dec(x_1); if (lean_obj_tag(x_107) == 0) { @@ -5431,20 +5332,6 @@ 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: { @@ -5453,16 +5340,17 @@ 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_8); 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) { +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___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) { _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); +x_7 = l_MonadExcept_ofExcept___at_Lean_Compiler_LCNF_PassManager_run___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); @@ -5479,6 +5367,7 @@ 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_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -5498,6 +5387,7 @@ x_18 = lean_unbox_float(x_8); lean_dec(x_8); x_19 = l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3(x_1, x_2, x_3, x_16, x_5, x_6, x_17, x_18, x_9, x_10, x_11, x_12, x_13, x_14, x_15); lean_dec(x_14); +lean_dec(x_13); lean_dec(x_11); lean_dec(x_5); return x_19; @@ -5539,11 +5429,11 @@ x_11 = l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2(x_ 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) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___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); +x_8 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___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); @@ -5552,7 +5442,7 @@ lean_dec(x_2); return x_8; } } -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) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___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, 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; @@ -5560,16 +5450,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__8(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__7(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__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_addTrace___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) { _start: { lean_object* x_8; -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); +x_8 = l_Lean_addTrace___at_Lean_Compiler_LCNF_PassManager_run___spec__8(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); @@ -5577,7 +5467,7 @@ lean_dec(x_3); return x_8; } } -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) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___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, 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; @@ -5585,12 +5475,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__10(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__9(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__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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___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) { _start: { size_t x_10; size_t x_11; lean_object* x_12; @@ -5598,14 +5488,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__11(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__10(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__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___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, lean_object* x_10) { _start: { size_t x_11; size_t x_12; lean_object* x_13; @@ -5613,7 +5503,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__12(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__11(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; } @@ -5905,7 +5795,7 @@ 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); +x_21 = l_Lean_profileitM___at_Lean_addDecl___spec__13___rarg(x_19, x_5, x_18, x_20, x_2, x_3, x_4); lean_dec(x_5); return x_21; } @@ -6394,52 +6284,52 @@ l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda_ lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__4___closed__2); l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__5___closed__1 = _init_l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__5___closed__1(); lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__5___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_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__1(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__1); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__2 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__2(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__2); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__3 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__3(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__3); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__4 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__4(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__4); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__5 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__5(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__5); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__6 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__6(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__6); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__7 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__7(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__7); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__8 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__8(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__8); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__9 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__9(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__9); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__10 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__10(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__10); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__11 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__11(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__11); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__12 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__12(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__12); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__13 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__13(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__13); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__14 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__14(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__14); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__15 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__15(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__15); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__16 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__16(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__16); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__17 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__17(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__17); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__18 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__18(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__18); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__19 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__19(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__7___lambda__1___closed__19); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__9___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__9___closed__1(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__9___closed__1); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__9___closed__2 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__9___closed__2(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__9___closed__2); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__9___closed__3 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__9___closed__3(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__9___closed__3); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__9___closed__4 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__9___closed__4(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__9___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(); diff --git a/stage0/stdlib/Lean/Compiler/Main.c b/stage0/stdlib/Lean/Compiler/Main.c index 494655b33049..7e3be5739ac3 100644 --- a/stage0/stdlib/Lean/Compiler/Main.c +++ b/stage0/stdlib/Lean/Compiler/Main.c @@ -32,11 +32,11 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_compile(lean_object*, lean_object*, lea static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_50____closed__11; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); 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_50____closed__9; static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_50____closed__16; static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_50____closed__18; static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_50____closed__3; +lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__13___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_50____closed__15; static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_50____closed__2; static lean_object* l_Lean_Compiler_compile___closed__1; @@ -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__14___rarg(x_7, x_5, x_6, x_8, x_2, x_3, x_4); +x_9 = l_Lean_profileitM___at_Lean_addDecl___spec__13___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/CoreM.c b/stage0/stdlib/Lean/CoreM.c index 0c11e5ca3956..4c227fb18579 100644 --- a/stage0/stdlib/Lean/CoreM.c +++ b/stage0/stdlib/Lean/CoreM.c @@ -13,7 +13,6 @@ #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*); @@ -22,6 +21,7 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Core_insta 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_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_addDecl___spec__12(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_39____closed__8; 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; @@ -84,6 +84,7 @@ 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*); size_t lean_uint64_to_usize(uint64_t); LEAN_EXPORT lean_object* l_Lean_Core_CoreM_run(lean_object*); +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_addDecl___spec__12___boxed(lean_object*, lean_object*, lean_object*, 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*); @@ -119,7 +120,6 @@ static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__ 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_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*); static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_39____closed__5; lean_object* l_Lean_Syntax_getPos_x3f(lean_object*, uint8_t); lean_object* lean_array_fget(lean_object*, lean_object*); @@ -138,6 +138,7 @@ LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUn lean_object* l_ReaderT_instMonadReaderT___rarg(lean_object*); static lean_object* l_Lean_Core_instMonadEnvCoreM___closed__2; uint8_t lean_float_decLt(double, double); +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__13(lean_object*); 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*); @@ -204,7 +205,6 @@ 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_mapCoreM___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -219,7 +219,6 @@ LEAN_EXPORT lean_object* l_Lean_Core_withCatchingRuntimeEx___rarg(uint8_t, lean_ 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_Lean_withoutCatchingRuntimeEx___rarg___lambda__1(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*); @@ -348,7 +347,6 @@ 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_checkInterrupted___boxed(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*); @@ -376,6 +374,7 @@ LEAN_EXPORT lean_object* l_Lean_withCatchingRuntimeEx(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*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadLogCoreM___lambda__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_instantiateValueLevelParams___lambda__2___closed__2; +lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNode___spec__1(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadQuotationCoreM; LEAN_EXPORT lean_object* l_Lean_Core_instantiateTypeLevelParams___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadLogCoreM___lambda__2(lean_object*, lean_object*, lean_object*); @@ -464,7 +463,6 @@ 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*); 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; @@ -530,7 +528,6 @@ LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_Core_instMeta 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; @@ -542,7 +539,6 @@ 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*); @@ -558,7 +554,9 @@ uint8_t l___private_Lean_Message_0__Lean_beqMessageSeverity____x40_Lean_Message_ uint8_t lean_level_eq(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_profileitM___at_Lean_addDecl___spec__13___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instInhabitedCache; +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__13___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_instMonadCoreM___closed__7; static lean_object* l_Lean_Core_instMonadQuotationCoreM___closed__2; LEAN_EXPORT lean_object* l_Lean_addAndCompile(lean_object*, lean_object*, lean_object*, lean_object*); @@ -10383,125 +10381,6 @@ 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: { @@ -10509,76 +10388,269 @@ 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; +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; 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; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; 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_12 = lean_st_ref_get(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); +x_15 = lean_ctor_get(x_13, 3); +lean_inc(x_15); +lean_dec(x_13); +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_addTraceNode___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 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_21, x_6, x_7, x_14); lean_dec(x_6); -return x_15; +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_st_ref_take(x_7, x_24); +x_26 = lean_ctor_get(x_25, 0); +lean_inc(x_26); +x_27 = lean_ctor_get(x_25, 1); +lean_inc(x_27); +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; uint8_t x_33; +x_29 = lean_ctor_get(x_26, 3); +lean_dec(x_29); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_3); +lean_ctor_set(x_30, 1, x_23); +x_31 = l_Lean_PersistentArray_push___rarg(x_1, x_30); +lean_ctor_set(x_26, 3, x_31); +x_32 = lean_st_ref_set(x_7, x_26, x_27); +x_33 = !lean_is_exclusive(x_32); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; +x_34 = lean_ctor_get(x_32, 0); +lean_dec(x_34); +x_35 = lean_box(0); +lean_ctor_set(x_32, 0, x_35); +return x_32; } 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; uint8_t x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; -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); -x_27 = lean_ctor_get_uint8(x_6, sizeof(void*)*11); -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_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_32, 1); +lean_inc(x_36); +lean_dec(x_32); +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_36); +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; lean_object* x_51; lean_object* x_52; +x_39 = lean_ctor_get(x_26, 0); +x_40 = lean_ctor_get(x_26, 1); +x_41 = lean_ctor_get(x_26, 2); +x_42 = lean_ctor_get(x_26, 4); +x_43 = lean_ctor_get(x_26, 5); +x_44 = lean_ctor_get(x_26, 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_dec(x_26); +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_3); +lean_ctor_set(x_45, 1, x_23); +x_46 = l_Lean_PersistentArray_push___rarg(x_1, x_45); +x_47 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_47, 0, x_39); +lean_ctor_set(x_47, 1, x_40); +lean_ctor_set(x_47, 2, x_41); +lean_ctor_set(x_47, 3, x_46); +lean_ctor_set(x_47, 4, x_42); +lean_ctor_set(x_47, 5, x_43); +lean_ctor_set(x_47, 6, x_44); +x_48 = lean_st_ref_set(x_7, x_47, x_27); +x_49 = lean_ctor_get(x_48, 1); +lean_inc(x_49); +if (lean_is_exclusive(x_48)) { + lean_ctor_release(x_48, 0); + lean_ctor_release(x_48, 1); + x_50 = x_48; +} else { + lean_dec_ref(x_48); + x_50 = lean_box(0); +} +x_51 = lean_box(0); +if (lean_is_scalar(x_50)) { + x_52 = lean_alloc_ctor(0, 2, 0); +} else { + x_52 = x_50; +} +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_49); +return x_52; +} +} +else +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; uint8_t x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; size_t x_73; size_t x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; +x_53 = lean_ctor_get(x_6, 0); +x_54 = lean_ctor_get(x_6, 1); +x_55 = lean_ctor_get(x_6, 2); +x_56 = lean_ctor_get(x_6, 3); +x_57 = lean_ctor_get(x_6, 4); +x_58 = lean_ctor_get(x_6, 5); +x_59 = lean_ctor_get(x_6, 6); +x_60 = lean_ctor_get(x_6, 7); +x_61 = lean_ctor_get(x_6, 8); +x_62 = lean_ctor_get(x_6, 9); +x_63 = lean_ctor_get(x_6, 10); +x_64 = lean_ctor_get_uint8(x_6, sizeof(void*)*11); +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_inc(x_55); +lean_inc(x_54); +lean_inc(x_53); lean_dec(x_6); -x_28 = l_Lean_replaceRef(x_3, x_21); -lean_dec(x_21); -x_29 = lean_alloc_ctor(0, 11, 1); -lean_ctor_set(x_29, 0, x_16); -lean_ctor_set(x_29, 1, x_17); -lean_ctor_set(x_29, 2, x_18); -lean_ctor_set(x_29, 3, x_19); -lean_ctor_set(x_29, 4, x_20); -lean_ctor_set(x_29, 5, x_28); -lean_ctor_set(x_29, 6, x_22); -lean_ctor_set(x_29, 7, x_23); -lean_ctor_set(x_29, 8, x_24); -lean_ctor_set(x_29, 9, x_25); -lean_ctor_set(x_29, 10, x_26); -lean_ctor_set_uint8(x_29, sizeof(void*)*11, x_27); -x_30 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_4, x_29, x_7, x_8); -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___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_addDecl___spec__12(x_1, x_2, x_3, x_31, x_5, x_29, x_7, x_32); -lean_dec(x_29); -return x_33; +x_65 = l_Lean_replaceRef(x_3, x_58); +lean_dec(x_58); +x_66 = lean_alloc_ctor(0, 11, 1); +lean_ctor_set(x_66, 0, x_53); +lean_ctor_set(x_66, 1, x_54); +lean_ctor_set(x_66, 2, x_55); +lean_ctor_set(x_66, 3, x_56); +lean_ctor_set(x_66, 4, x_57); +lean_ctor_set(x_66, 5, x_65); +lean_ctor_set(x_66, 6, x_59); +lean_ctor_set(x_66, 7, x_60); +lean_ctor_set(x_66, 8, x_61); +lean_ctor_set(x_66, 9, x_62); +lean_ctor_set(x_66, 10, x_63); +lean_ctor_set_uint8(x_66, sizeof(void*)*11, x_64); +x_67 = lean_st_ref_get(x_7, x_8); +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_68, 3); +lean_inc(x_70); +lean_dec(x_68); +x_71 = l_Lean_PersistentArray_toArray___rarg(x_70); +x_72 = lean_array_get_size(x_71); +x_73 = lean_usize_of_nat(x_72); +lean_dec(x_72); +x_74 = 0; +x_75 = l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNode___spec__1(x_73, x_74, x_71); +x_76 = lean_alloc_ctor(9, 3, 1); +lean_ctor_set(x_76, 0, x_2); +lean_ctor_set(x_76, 1, x_4); +lean_ctor_set(x_76, 2, x_75); +lean_ctor_set_uint8(x_76, sizeof(void*)*3, x_5); +x_77 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_76, x_66, x_7, x_69); +lean_dec(x_66); +x_78 = lean_ctor_get(x_77, 0); +lean_inc(x_78); +x_79 = lean_ctor_get(x_77, 1); +lean_inc(x_79); +lean_dec(x_77); +x_80 = lean_st_ref_take(x_7, x_79); +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); +x_84 = lean_ctor_get(x_81, 1); +lean_inc(x_84); +x_85 = lean_ctor_get(x_81, 2); +lean_inc(x_85); +x_86 = lean_ctor_get(x_81, 4); +lean_inc(x_86); +x_87 = lean_ctor_get(x_81, 5); +lean_inc(x_87); +x_88 = lean_ctor_get(x_81, 6); +lean_inc(x_88); +if (lean_is_exclusive(x_81)) { + lean_ctor_release(x_81, 0); + lean_ctor_release(x_81, 1); + lean_ctor_release(x_81, 2); + lean_ctor_release(x_81, 3); + lean_ctor_release(x_81, 4); + lean_ctor_release(x_81, 5); + lean_ctor_release(x_81, 6); + x_89 = x_81; +} else { + lean_dec_ref(x_81); + x_89 = lean_box(0); +} +x_90 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_90, 0, x_3); +lean_ctor_set(x_90, 1, x_78); +x_91 = l_Lean_PersistentArray_push___rarg(x_1, x_90); +if (lean_is_scalar(x_89)) { + x_92 = lean_alloc_ctor(0, 7, 0); +} else { + x_92 = x_89; +} +lean_ctor_set(x_92, 0, x_83); +lean_ctor_set(x_92, 1, x_84); +lean_ctor_set(x_92, 2, x_85); +lean_ctor_set(x_92, 3, x_91); +lean_ctor_set(x_92, 4, x_86); +lean_ctor_set(x_92, 5, x_87); +lean_ctor_set(x_92, 6, x_88); +x_93 = lean_st_ref_set(x_7, x_92, x_82); +x_94 = lean_ctor_get(x_93, 1); +lean_inc(x_94); +if (lean_is_exclusive(x_93)) { + lean_ctor_release(x_93, 0); + lean_ctor_release(x_93, 1); + x_95 = x_93; +} else { + lean_dec_ref(x_93); + x_95 = lean_box(0); +} +x_96 = lean_box(0); +if (lean_is_scalar(x_95)) { + x_97 = lean_alloc_ctor(0, 2, 0); +} else { + x_97 = x_95; +} +lean_ctor_set(x_97, 0, x_96); +lean_ctor_set(x_97, 1, x_94); +return x_97; } } } -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) { +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_addDecl___spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { if (lean_obj_tag(x_1) == 0) @@ -10678,7 +10750,7 @@ x_11 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_addDecl___spec 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); +x_13 = l_MonadExcept_ofExcept___at_Lean_addDecl___spec__12(x_5, x_8, x_9, x_12); lean_dec(x_8); return x_13; } @@ -10936,7 +11008,7 @@ x_28 = lean_st_ref_set(x_10, x_23, x_24); x_29 = lean_ctor_get(x_28, 1); lean_inc(x_29); lean_dec(x_28); -x_30 = l_MonadExcept_ofExcept___at_Lean_addDecl___spec__13(x_19, x_9, x_10, x_29); +x_30 = l_MonadExcept_ofExcept___at_Lean_addDecl___spec__12(x_19, x_9, x_10, x_29); lean_dec(x_10); lean_dec(x_9); lean_dec(x_19); @@ -11016,7 +11088,7 @@ x_48 = lean_st_ref_set(x_10, x_47, x_24); x_49 = lean_ctor_get(x_48, 1); lean_inc(x_49); lean_dec(x_48); -x_50 = l_MonadExcept_ofExcept___at_Lean_addDecl___spec__13(x_19, x_9, x_10, x_49); +x_50 = l_MonadExcept_ofExcept___at_Lean_addDecl___spec__12(x_19, x_9, x_10, x_49); lean_dec(x_10); lean_dec(x_9); lean_dec(x_19); @@ -11217,7 +11289,7 @@ 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) { +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__13___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; @@ -11226,11 +11298,11 @@ 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) { +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__13(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); +x_2 = lean_alloc_closure((void*)(l_Lean_profileitM___at_Lean_addDecl___spec__13___rarg___boxed), 7, 0); return x_2; } } @@ -11416,7 +11488,7 @@ 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); +x_14 = l_Lean_profileitM___at_Lean_addDecl___spec__13___rarg(x_12, x_5, x_11, x_13, x_2, x_3, x_4); lean_dec(x_5); return x_14; } @@ -11501,18 +11573,6 @@ 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: { @@ -11524,11 +11584,11 @@ 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) { +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_addDecl___spec__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_MonadExcept_ofExcept___at_Lean_addDecl___spec__13(x_1, x_2, x_3, x_4); +x_5 = l_MonadExcept_ofExcept___at_Lean_addDecl___spec__12(x_1, x_2, x_3, x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); @@ -11600,11 +11660,11 @@ x_9 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7(x_1, x_2, x_3, x_8, x_5, 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) { +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__13___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__14___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +x_8 = l_Lean_profileitM___at_Lean_addDecl___spec__13___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; diff --git a/stage0/stdlib/Lean/Elab/Command.c b/stage0/stdlib/Lean/Elab/Command.c index 7110e21b3632..00c6862e616c 100644 --- a/stage0/stdlib/Lean/Elab/Command.c +++ b/stage0/stdlib/Lean/Elab/Command.c @@ -13,8 +13,9 @@ #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_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__16___boxed(lean_object*, lean_object*, lean_object*, 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_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__10(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_Context_cmdPos___default; lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__12___boxed(lean_object*, lean_object*); @@ -34,8 +35,10 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadQuotationCommandElabM___la 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_withTraceNode___at_Lean_Elab_Command_runLinters___spec__11___lambda__2(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, double, lean_object*, lean_object*, lean_object*, 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_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__13___rarg___boxed(lean_object*, lean_object*); lean_object* l_Lean_MapDeclarationExtension_insert___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_lintersRef; static lean_object* l_Lean_Elab_Command_getBracketedBinderIds___closed__5; @@ -50,16 +53,19 @@ static lean_object* l___auto____x40_Lean_Elab_Command___hyg_404____closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_addCompletionInfo___at_Lean_withSetOptionIn___spec__2(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*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Command_runLinters___spec__15(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabSetOption___at_Lean_withSetOptionIn___spec__1___closed__4; 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*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__10___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_shift_right(size_t, size_t); static lean_object* l_Lean_Elab_Command_instMonadRecDepthCommandElabM___closed__1; lean_object* l_Lean_Elab_CommandContextInfo_save___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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*); 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_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__12(lean_object*, lean_object*, lean_object*, 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*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_516_(lean_object*); @@ -123,11 +129,9 @@ static lean_object* l_Lean_Elab_expandDeclId___at_Lean_Elab_Command_expandDeclId 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*); 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*); 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*); @@ -146,7 +150,6 @@ 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_Elab_Command_instMonadMacroAdapterCommandElabM___lambda__1___boxed(lean_object*, lean_object*, lean_object*); @@ -155,7 +158,6 @@ 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_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2939_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabSetOption_setOption___at_Lean_withSetOptionIn___spec__4___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -196,7 +198,6 @@ LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at_Lean_Elab_Command_ex LEAN_EXPORT lean_object* l_Lean_Elab_Command_getScopes___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at_Lean_Elab_Command_liftTermElabM___spec__13(lean_object*, lean_object*); 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*); static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_548____closed__4; static lean_object* l_Lean_Elab_Command_elabCommandTopLevel___lambda__2___closed__1; 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*); @@ -227,10 +228,8 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_runLinters___lambda__1___boxed(lean 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*, 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*); static lean_object* l___auto____x40_Lean_Elab_Command___hyg_404____closed__9; 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*); static lean_object* l___auto____x40_Lean_Elab_Command___hyg_404____closed__11; 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*); @@ -240,6 +239,7 @@ 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*); +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Command_runLinters___spec__17___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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*); @@ -250,14 +250,15 @@ uint8_t l___private_Lean_Util_Trace_0__Lean_checkTraceOption(lean_object*, lean_ LEAN_EXPORT lean_object* l_Lean_withSetOptionIn(lean_object*, lean_object*, lean_object*, lean_object*, 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*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__10___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_maxRecDepth; lean_object* l_Lean_KVMap_insertCore(lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__10___lambda__1___closed__2; lean_object* l_List_head_x21___rarg(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*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at_Lean_Elab_Command_liftTermElabM___spec__13___boxed(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; @@ -331,13 +332,12 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCo 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_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*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__13(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_liftCommandElabM___spec__1___boxed(lean_object*, lean_object*, lean_object*, 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; @@ -348,6 +348,7 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runTerm 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_withSeconds___at_Lean_Elab_Command_runLinters___spec__14(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; @@ -378,6 +379,7 @@ LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at___private_Lean_Elab_Comman 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_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__10___lambda__1___closed__1; 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; @@ -399,7 +401,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadLiftTIOCommandElabM(lean_o static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Command_liftTermElabM___spec__1___rarg___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_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*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_withSetOptionIn___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__7; static lean_object* l___auto____x40_Lean_Elab_Command___hyg_404____closed__2; @@ -410,17 +411,18 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_getRef(lean_object*, lean_object*, 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_withTraceNode___at_Lean_Elab_Command_runLinters___spec__11___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_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*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__13___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___spec__7(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Command_runLinters___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getKind(lean_object*); static lean_object* l_Lean_withSetOptionIn___closed__1; static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___lambda__1___closed__2; @@ -433,11 +435,9 @@ 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_Lean_liftCommandElabM(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_PersistentArray_mapM___at_Lean_Elab_Command_liftCoreM___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___spec__12(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_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12___lambda__3(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, double, lean_object*, lean_object*, 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*); 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; @@ -455,8 +455,10 @@ 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*); LEAN_EXPORT lean_object* l_Lean_liftCommandElabM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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_Lean_profileitM___at_Lean_Elab_Command_runLinters___spec__17(lean_object*); static lean_object* l___auto____x40_Lean_Elab_Command___hyg_404____closed__26; 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*); @@ -475,10 +477,10 @@ uint8_t l_Lean_MessageData_hasTag(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at_Lean_Elab_Command_liftTermElabM___spec__9(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*); static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__4___closed__2; -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_withSetOptionIn___closed__4; static lean_object* l_Lean_Elab_Command_instMonadTraceCommandElabM___closed__1; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___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*); static lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Command_expandDeclId___spec__7___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Command_State_ngen___default; static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_548____closed__5; @@ -503,9 +505,9 @@ static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_54 static lean_object* l_Lean_Elab_Command_instMonadRefCommandElabM___closed__1; static lean_object* l_Lean_addMessageContextPartial___at_Lean_Elab_Command_instAddMessageContextCommandElabM___spec__1___closed__2; static lean_object* l_Lean_Elab_Command_liftTermElabM___rarg___closed__16; -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12___lambda__4(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__5___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_getMainModule___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___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* 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*); @@ -540,6 +542,7 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_expandDeclId__ LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommand___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_warningAsError; extern lean_object* l_Lean_Elab_pp_macroStack; +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Command_runLinters___spec__17___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_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*); static lean_object* l___auto____x40_Lean_Elab_Command___hyg_404____closed__3; @@ -556,12 +559,13 @@ lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_ static lean_object* l_Lean_Elab_Command_instMonadQuotationCommandElabM___closed__4; static lean_object* l_Lean_liftCommandElabM___rarg___closed__2; 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*, uint8_t, lean_object*, lean_object*, uint8_t, double, lean_object*, 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__10___lambda__2(lean_object*, 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; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__11___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_modifyScope___closed__4; static lean_object* l_Lean_Elab_Command_liftTermElabM___rarg___closed__2; static lean_object* l___auto____x40_Lean_Elab_Command___hyg_404____closed__12; @@ -576,7 +580,6 @@ static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandT 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; @@ -584,27 +587,26 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermE LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadLogCommandElabM___lambda__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addDocString_x27___at_Lean_Elab_Command_expandDeclId___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withSetOptionIn___lambda__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__11___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_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*); 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_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__11(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2971____closed__2; -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_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_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_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12___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_Elab_Command_addLinter(lean_object*, lean_object*); lean_object* lean_usize_to_nat(size_t); size_t lean_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*); 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___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*); 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*); @@ -617,6 +619,7 @@ LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabCo LEAN_EXPORT lean_object* l_Lean_Elab_Command_runTermElabM(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadOptionsCommandElabM___boxed(lean_object*); static lean_object* l___auto____x40_Lean_Elab_Command___hyg_404____closed__8; +lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNode___spec__1(size_t, size_t, 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*); @@ -658,7 +661,6 @@ static lean_object* l_Lean_Elab_Command_elabCommand___closed__3; static lean_object* l_Lean_Elab_Command_instMonadInfoTreeCommandElabM___closed__3; static lean_object* l_Lean_Elab_elabSetOption___at_Lean_withSetOptionIn___spec__1___closed__2; 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; @@ -667,7 +669,8 @@ lean_object* l_Lean_Syntax_setArgs(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadOptionsCommandElabM___rarg___boxed(lean_object*, lean_object*); 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_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__16(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__11___lambda__4(lean_object*, 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_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*); @@ -681,7 +684,6 @@ static lean_object* l_Lean_Elab_Command_instMonadLogCommandElabM___closed__5; static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_548____closed__16; static lean_object* l_Lean_Elab_throwAlreadyDeclaredUniverseLevel___at_Lean_Elab_Command_addUnivLevel___spec__1___closed__4; lean_object* l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_withDeclName___spec__3___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Command_runLinters___spec__19(lean_object*); static lean_object* l_Lean_Elab_elabSetOption___at_Lean_withSetOptionIn___spec__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Command_getScope(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadLogCommandElabM___lambda__3(lean_object*, lean_object*, lean_object*); @@ -697,7 +699,6 @@ 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_object*); 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_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_548____closed__3; -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_Elab_Command_adaptExpander(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -724,7 +725,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadTraceCommandElabM___lambda lean_object* l_Lean_TagDeclarationExtension_tag(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*); static lean_object* l_Lean_Elab_elabSetOption_setOption___at_Lean_withSetOptionIn___spec__4___closed__1; -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_elabSetOption_setOption___at_Lean_withSetOptionIn___spec__4___boxed(lean_object*, lean_object*, 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*); @@ -736,7 +736,6 @@ static lean_object* l___auto____x40_Lean_Elab_Command___hyg_404____closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at_Lean_Elab_Command_liftTermElabM___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_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; uint8_t l_Lean_Syntax_isNone(lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_Command_modifyScope___spec__1(lean_object*); @@ -776,7 +775,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_throwAlreadyDeclaredUniverseLevel___at_Lean 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_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*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Command_liftTermElabM___spec__18___rarg(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; @@ -784,12 +782,12 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCom LEAN_EXPORT lean_object* l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSavedPartialInfoContext___at_Lean_Elab_Command_liftTermElabM___spec__2(lean_object*); static lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__1___closed__5; lean_object* l_Lean_getOptionDecl(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__11___lambda__3(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_548____closed__17; lean_object* l_Array_ofSubarray___rarg(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_object* lean_erase_macro_scopes(lean_object*); @@ -820,7 +818,6 @@ static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId 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_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); @@ -830,7 +827,9 @@ 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; +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__13___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_ioErrorToMessage___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); 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*); @@ -839,6 +838,7 @@ uint8_t l_Lean_PersistentArray_isEmpty___rarg(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; +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__9(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_io_error_to_string(lean_object*); static lean_object* l___auto____x40_Lean_Elab_Command___hyg_404____closed__10; lean_object* l_Lean_KeyedDeclsAttribute_getEntries___rarg(lean_object*, lean_object*, lean_object*); @@ -892,16 +892,13 @@ static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_29 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_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*); static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_548____closed__11; static lean_object* l___auto____x40_Lean_Elab_Command___hyg_404____closed__19; -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___auto____x40_Lean_Elab_Command___hyg_404____closed__21; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__5___closed__4; @@ -914,6 +911,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_getResetInfoTrees___at___private_Lean_Elab_ 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_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__9___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__6; static lean_object* l___auto____x40_Lean_Elab_Command___hyg_404____closed__18; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___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*); @@ -927,10 +925,10 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkTe 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; lean_object* l_Lean_getStructureFieldsFlattened(lean_object*, lean_object*, uint8_t); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_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*); @@ -953,7 +951,6 @@ static lean_object* l___auto____x40_Lean_Elab_Command___hyg_404____closed__6; 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*); static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_548____closed__10; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__5(lean_object*, 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_Lean_Elab_Command_instMonadRecDepthCommandElabM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -966,7 +963,6 @@ LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___private_Lean_Elab_C LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftTermElabM___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_pure___at_Lean_Elab_liftMacroM___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*); 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*); LEAN_EXPORT lean_object* l_Lean_liftCommandElabM___rarg___lambda__1(lean_object*); 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; @@ -8767,131 +8763,6 @@ 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_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; 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__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: { @@ -8907,60 +8778,267 @@ 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; +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; size_t x_21; size_t x_22; lean_object* x_23; lean_object* x_24; 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_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_15 = lean_st_ref_get(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); +x_18 = lean_ctor_get(x_16, 8); +lean_inc(x_18); +lean_dec(x_16); +x_19 = l_Lean_PersistentArray_toArray___rarg(x_18); +x_20 = lean_array_get_size(x_19); +x_21 = lean_usize_of_nat(x_20); +lean_dec(x_20); +x_22 = 0; +x_23 = l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNode___spec__1(x_21, x_22, x_19); +x_24 = lean_alloc_ctor(9, 3, 1); +lean_ctor_set(x_24, 0, x_2); +lean_ctor_set(x_24, 1, x_4); +lean_ctor_set(x_24, 2, x_23); +lean_ctor_set_uint8(x_24, sizeof(void*)*3, x_5); +x_25 = l_Lean_addMessageContextPartial___at_Lean_Elab_Command_instAddMessageContextCommandElabM___spec__1(x_24, x_6, x_7, x_17); lean_dec(x_6); -return x_18; +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_st_ref_take(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 = !lean_is_exclusive(x_29); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_32 = lean_ctor_get(x_29, 8); +lean_dec(x_32); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_3); +lean_ctor_set(x_33, 1, x_26); +x_34 = l_Lean_PersistentArray_push___rarg(x_1, x_33); +lean_ctor_set(x_29, 8, x_34); +x_35 = lean_st_ref_set(x_7, x_29, x_30); +x_36 = !lean_is_exclusive(x_35); +if (x_36 == 0) +{ +lean_object* x_37; lean_object* x_38; +x_37 = lean_ctor_get(x_35, 0); +lean_dec(x_37); +x_38 = lean_box(0); +lean_ctor_set(x_35, 0, x_38); +return x_35; } else { -lean_object* x_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_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_35, 1); +lean_inc(x_39); +lean_dec(x_35); +x_40 = lean_box(0); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_39); +return x_41; +} +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; 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_42 = lean_ctor_get(x_29, 0); +x_43 = lean_ctor_get(x_29, 1); +x_44 = lean_ctor_get(x_29, 2); +x_45 = lean_ctor_get(x_29, 3); +x_46 = lean_ctor_get(x_29, 4); +x_47 = lean_ctor_get(x_29, 5); +x_48 = lean_ctor_get(x_29, 6); +x_49 = lean_ctor_get(x_29, 7); +lean_inc(x_49); +lean_inc(x_48); +lean_inc(x_47); +lean_inc(x_46); +lean_inc(x_45); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_29); +x_50 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_50, 0, x_3); +lean_ctor_set(x_50, 1, x_26); +x_51 = l_Lean_PersistentArray_push___rarg(x_1, x_50); +x_52 = lean_alloc_ctor(0, 9, 0); +lean_ctor_set(x_52, 0, x_42); +lean_ctor_set(x_52, 1, x_43); +lean_ctor_set(x_52, 2, x_44); +lean_ctor_set(x_52, 3, x_45); +lean_ctor_set(x_52, 4, x_46); +lean_ctor_set(x_52, 5, x_47); +lean_ctor_set(x_52, 6, x_48); +lean_ctor_set(x_52, 7, x_49); +lean_ctor_set(x_52, 8, x_51); +x_53 = lean_st_ref_set(x_7, x_52, x_30); +x_54 = lean_ctor_get(x_53, 1); +lean_inc(x_54); +if (lean_is_exclusive(x_53)) { + lean_ctor_release(x_53, 0); + lean_ctor_release(x_53, 1); + x_55 = x_53; +} else { + lean_dec_ref(x_53); + x_55 = lean_box(0); +} +x_56 = lean_box(0); +if (lean_is_scalar(x_55)) { + x_57 = lean_alloc_ctor(0, 2, 0); +} else { + x_57 = x_55; +} +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_54); +return x_57; +} +} +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; size_t x_72; size_t x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_58 = lean_ctor_get(x_6, 0); +x_59 = lean_ctor_get(x_6, 1); +x_60 = lean_ctor_get(x_6, 2); +x_61 = lean_ctor_get(x_6, 3); +x_62 = lean_ctor_get(x_6, 4); +x_63 = lean_ctor_get(x_6, 5); +x_64 = lean_ctor_get(x_6, 7); +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_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; +x_65 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_65, 0, x_58); +lean_ctor_set(x_65, 1, x_59); +lean_ctor_set(x_65, 2, x_60); +lean_ctor_set(x_65, 3, x_61); +lean_ctor_set(x_65, 4, x_62); +lean_ctor_set(x_65, 5, x_63); +lean_ctor_set(x_65, 6, x_12); +lean_ctor_set(x_65, 7, x_64); +x_66 = lean_st_ref_get(x_7, x_11); +x_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_66, 1); +lean_inc(x_68); +lean_dec(x_66); +x_69 = lean_ctor_get(x_67, 8); +lean_inc(x_69); +lean_dec(x_67); +x_70 = l_Lean_PersistentArray_toArray___rarg(x_69); +x_71 = lean_array_get_size(x_70); +x_72 = lean_usize_of_nat(x_71); +lean_dec(x_71); +x_73 = 0; +x_74 = l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNode___spec__1(x_72, x_73, x_70); +x_75 = lean_alloc_ctor(9, 3, 1); +lean_ctor_set(x_75, 0, x_2); +lean_ctor_set(x_75, 1, x_4); +lean_ctor_set(x_75, 2, x_74); +lean_ctor_set_uint8(x_75, sizeof(void*)*3, x_5); +x_76 = l_Lean_addMessageContextPartial___at_Lean_Elab_Command_instAddMessageContextCommandElabM___spec__1(x_75, x_65, x_7, x_68); +lean_dec(x_65); +x_77 = lean_ctor_get(x_76, 0); +lean_inc(x_77); +x_78 = lean_ctor_get(x_76, 1); +lean_inc(x_78); +lean_dec(x_76); +x_79 = lean_st_ref_take(x_7, 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 = lean_ctor_get(x_80, 0); +lean_inc(x_82); +x_83 = lean_ctor_get(x_80, 1); +lean_inc(x_83); +x_84 = lean_ctor_get(x_80, 2); +lean_inc(x_84); +x_85 = lean_ctor_get(x_80, 3); +lean_inc(x_85); +x_86 = lean_ctor_get(x_80, 4); +lean_inc(x_86); +x_87 = lean_ctor_get(x_80, 5); +lean_inc(x_87); +x_88 = lean_ctor_get(x_80, 6); +lean_inc(x_88); +x_89 = lean_ctor_get(x_80, 7); +lean_inc(x_89); +if (lean_is_exclusive(x_80)) { + lean_ctor_release(x_80, 0); + lean_ctor_release(x_80, 1); + lean_ctor_release(x_80, 2); + lean_ctor_release(x_80, 3); + lean_ctor_release(x_80, 4); + lean_ctor_release(x_80, 5); + lean_ctor_release(x_80, 6); + lean_ctor_release(x_80, 7); + lean_ctor_release(x_80, 8); + x_90 = x_80; +} else { + lean_dec_ref(x_80); + x_90 = lean_box(0); +} +x_91 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_91, 0, x_3); +lean_ctor_set(x_91, 1, x_77); +x_92 = l_Lean_PersistentArray_push___rarg(x_1, x_91); +if (lean_is_scalar(x_90)) { + x_93 = lean_alloc_ctor(0, 9, 0); +} else { + x_93 = x_90; +} +lean_ctor_set(x_93, 0, x_82); +lean_ctor_set(x_93, 1, x_83); +lean_ctor_set(x_93, 2, x_84); +lean_ctor_set(x_93, 3, x_85); +lean_ctor_set(x_93, 4, x_86); +lean_ctor_set(x_93, 5, x_87); +lean_ctor_set(x_93, 6, x_88); +lean_ctor_set(x_93, 7, x_89); +lean_ctor_set(x_93, 8, x_92); +x_94 = lean_st_ref_set(x_7, x_93, x_81); +x_95 = lean_ctor_get(x_94, 1); +lean_inc(x_95); +if (lean_is_exclusive(x_94)) { + lean_ctor_release(x_94, 0); + lean_ctor_release(x_94, 1); + x_96 = x_94; +} else { + lean_dec_ref(x_94); + x_96 = lean_box(0); +} +x_97 = lean_box(0); +if (lean_is_scalar(x_96)) { + x_98 = lean_alloc_ctor(0, 2, 0); +} else { + x_98 = x_96; +} +lean_ctor_set(x_98, 0, x_97); +lean_ctor_set(x_98, 1, x_95); +return x_98; } } } -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) { +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { if (lean_obj_tag(x_1) == 0) @@ -9060,7 +9138,7 @@ x_11 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Command_r 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); +x_13 = l_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__9(x_5, x_8, x_9, x_12); lean_dec(x_8); return x_13; } @@ -9322,7 +9400,7 @@ x_28 = lean_st_ref_set(x_10, x_23, x_24); x_29 = lean_ctor_get(x_28, 1); lean_inc(x_29); lean_dec(x_28); -x_30 = l_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__10(x_19, x_9, x_10, x_29); +x_30 = l_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__9(x_19, x_9, x_10, x_29); lean_dec(x_10); lean_dec(x_9); lean_dec(x_19); @@ -9408,7 +9486,7 @@ x_50 = lean_st_ref_set(x_10, x_49, x_24); x_51 = lean_ctor_get(x_50, 1); lean_inc(x_51); lean_dec(x_50); -x_52 = l_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__10(x_19, x_9, x_10, x_51); +x_52 = l_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__9(x_19, x_9, x_10, x_51); lean_dec(x_10); lean_dec(x_9); lean_dec(x_19); @@ -9622,7 +9700,7 @@ return x_36; } } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__11___lambda__1___closed__1() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__10___lambda__1___closed__1() { _start: { lean_object* x_1; @@ -9630,16 +9708,16 @@ 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() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__10___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_1 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__10___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) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__10___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; @@ -9647,7 +9725,7 @@ 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_8 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__10___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); @@ -9661,7 +9739,7 @@ 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) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__10___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; @@ -10053,7 +10131,7 @@ 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) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { uint8_t x_10; @@ -10076,10 +10154,10 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_obje 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); +x_13 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__10___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); +x_14 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__10___lambda__2), 5, 2); lean_closure_set(x_14, 0, x_12); lean_closure_set(x_14, 1, x_1); x_15 = 1; @@ -10130,7 +10208,7 @@ 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) { +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; @@ -10191,7 +10269,7 @@ 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) { +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__13___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; @@ -10293,15 +10371,15 @@ 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) { +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__13(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); +x_2 = lean_alloc_closure((void*)(l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__13___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) { +LEAN_EXPORT lean_object* l_Lean_withSeconds___at_Lean_Elab_Command_runLinters___spec__14(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; @@ -10394,132 +10472,7 @@ 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) { +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Command_runLinters___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) { _start: { lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; @@ -10534,60 +10487,267 @@ 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; +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; size_t x_21; size_t x_22; lean_object* x_23; lean_object* x_24; 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_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_15 = lean_st_ref_get(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); +x_18 = lean_ctor_get(x_16, 8); +lean_inc(x_18); +lean_dec(x_16); +x_19 = l_Lean_PersistentArray_toArray___rarg(x_18); +x_20 = lean_array_get_size(x_19); +x_21 = lean_usize_of_nat(x_20); +lean_dec(x_20); +x_22 = 0; +x_23 = l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNode___spec__1(x_21, x_22, x_19); +x_24 = lean_alloc_ctor(9, 3, 1); +lean_ctor_set(x_24, 0, x_2); +lean_ctor_set(x_24, 1, x_4); +lean_ctor_set(x_24, 2, x_23); +lean_ctor_set_uint8(x_24, sizeof(void*)*3, x_5); +x_25 = l_Lean_addMessageContextPartial___at_Lean_Elab_Command_instAddMessageContextCommandElabM___spec__1(x_24, x_6, x_7, x_17); lean_dec(x_6); -return x_18; +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_st_ref_take(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 = !lean_is_exclusive(x_29); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_32 = lean_ctor_get(x_29, 8); +lean_dec(x_32); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_3); +lean_ctor_set(x_33, 1, x_26); +x_34 = l_Lean_PersistentArray_push___rarg(x_1, x_33); +lean_ctor_set(x_29, 8, x_34); +x_35 = lean_st_ref_set(x_7, x_29, x_30); +x_36 = !lean_is_exclusive(x_35); +if (x_36 == 0) +{ +lean_object* x_37; lean_object* x_38; +x_37 = lean_ctor_get(x_35, 0); +lean_dec(x_37); +x_38 = lean_box(0); +lean_ctor_set(x_35, 0, x_38); +return x_35; } else { -lean_object* x_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_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_35, 1); +lean_inc(x_39); +lean_dec(x_35); +x_40 = lean_box(0); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_39); +return x_41; +} +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; 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_42 = lean_ctor_get(x_29, 0); +x_43 = lean_ctor_get(x_29, 1); +x_44 = lean_ctor_get(x_29, 2); +x_45 = lean_ctor_get(x_29, 3); +x_46 = lean_ctor_get(x_29, 4); +x_47 = lean_ctor_get(x_29, 5); +x_48 = lean_ctor_get(x_29, 6); +x_49 = lean_ctor_get(x_29, 7); +lean_inc(x_49); +lean_inc(x_48); +lean_inc(x_47); +lean_inc(x_46); +lean_inc(x_45); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_29); +x_50 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_50, 0, x_3); +lean_ctor_set(x_50, 1, x_26); +x_51 = l_Lean_PersistentArray_push___rarg(x_1, x_50); +x_52 = lean_alloc_ctor(0, 9, 0); +lean_ctor_set(x_52, 0, x_42); +lean_ctor_set(x_52, 1, x_43); +lean_ctor_set(x_52, 2, x_44); +lean_ctor_set(x_52, 3, x_45); +lean_ctor_set(x_52, 4, x_46); +lean_ctor_set(x_52, 5, x_47); +lean_ctor_set(x_52, 6, x_48); +lean_ctor_set(x_52, 7, x_49); +lean_ctor_set(x_52, 8, x_51); +x_53 = lean_st_ref_set(x_7, x_52, x_30); +x_54 = lean_ctor_get(x_53, 1); +lean_inc(x_54); +if (lean_is_exclusive(x_53)) { + lean_ctor_release(x_53, 0); + lean_ctor_release(x_53, 1); + x_55 = x_53; +} else { + lean_dec_ref(x_53); + x_55 = lean_box(0); +} +x_56 = lean_box(0); +if (lean_is_scalar(x_55)) { + x_57 = lean_alloc_ctor(0, 2, 0); +} else { + x_57 = x_55; +} +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_54); +return x_57; +} +} +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; size_t x_72; size_t x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_58 = lean_ctor_get(x_6, 0); +x_59 = lean_ctor_get(x_6, 1); +x_60 = lean_ctor_get(x_6, 2); +x_61 = lean_ctor_get(x_6, 3); +x_62 = lean_ctor_get(x_6, 4); +x_63 = lean_ctor_get(x_6, 5); +x_64 = lean_ctor_get(x_6, 7); +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_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; +x_65 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_65, 0, x_58); +lean_ctor_set(x_65, 1, x_59); +lean_ctor_set(x_65, 2, x_60); +lean_ctor_set(x_65, 3, x_61); +lean_ctor_set(x_65, 4, x_62); +lean_ctor_set(x_65, 5, x_63); +lean_ctor_set(x_65, 6, x_12); +lean_ctor_set(x_65, 7, x_64); +x_66 = lean_st_ref_get(x_7, x_11); +x_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_66, 1); +lean_inc(x_68); +lean_dec(x_66); +x_69 = lean_ctor_get(x_67, 8); +lean_inc(x_69); +lean_dec(x_67); +x_70 = l_Lean_PersistentArray_toArray___rarg(x_69); +x_71 = lean_array_get_size(x_70); +x_72 = lean_usize_of_nat(x_71); +lean_dec(x_71); +x_73 = 0; +x_74 = l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNode___spec__1(x_72, x_73, x_70); +x_75 = lean_alloc_ctor(9, 3, 1); +lean_ctor_set(x_75, 0, x_2); +lean_ctor_set(x_75, 1, x_4); +lean_ctor_set(x_75, 2, x_74); +lean_ctor_set_uint8(x_75, sizeof(void*)*3, x_5); +x_76 = l_Lean_addMessageContextPartial___at_Lean_Elab_Command_instAddMessageContextCommandElabM___spec__1(x_75, x_65, x_7, x_68); +lean_dec(x_65); +x_77 = lean_ctor_get(x_76, 0); +lean_inc(x_77); +x_78 = lean_ctor_get(x_76, 1); +lean_inc(x_78); +lean_dec(x_76); +x_79 = lean_st_ref_take(x_7, 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 = lean_ctor_get(x_80, 0); +lean_inc(x_82); +x_83 = lean_ctor_get(x_80, 1); +lean_inc(x_83); +x_84 = lean_ctor_get(x_80, 2); +lean_inc(x_84); +x_85 = lean_ctor_get(x_80, 3); +lean_inc(x_85); +x_86 = lean_ctor_get(x_80, 4); +lean_inc(x_86); +x_87 = lean_ctor_get(x_80, 5); +lean_inc(x_87); +x_88 = lean_ctor_get(x_80, 6); +lean_inc(x_88); +x_89 = lean_ctor_get(x_80, 7); +lean_inc(x_89); +if (lean_is_exclusive(x_80)) { + lean_ctor_release(x_80, 0); + lean_ctor_release(x_80, 1); + lean_ctor_release(x_80, 2); + lean_ctor_release(x_80, 3); + lean_ctor_release(x_80, 4); + lean_ctor_release(x_80, 5); + lean_ctor_release(x_80, 6); + lean_ctor_release(x_80, 7); + lean_ctor_release(x_80, 8); + x_90 = x_80; +} else { + lean_dec_ref(x_80); + x_90 = lean_box(0); +} +x_91 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_91, 0, x_3); +lean_ctor_set(x_91, 1, x_77); +x_92 = l_Lean_PersistentArray_push___rarg(x_1, x_91); +if (lean_is_scalar(x_90)) { + x_93 = lean_alloc_ctor(0, 9, 0); +} else { + x_93 = x_90; +} +lean_ctor_set(x_93, 0, x_82); +lean_ctor_set(x_93, 1, x_83); +lean_ctor_set(x_93, 2, x_84); +lean_ctor_set(x_93, 3, x_85); +lean_ctor_set(x_93, 4, x_86); +lean_ctor_set(x_93, 5, x_87); +lean_ctor_set(x_93, 6, x_88); +lean_ctor_set(x_93, 7, x_89); +lean_ctor_set(x_93, 8, x_92); +x_94 = lean_st_ref_set(x_7, x_93, x_81); +x_95 = lean_ctor_get(x_94, 1); +lean_inc(x_95); +if (lean_is_exclusive(x_94)) { + lean_ctor_release(x_94, 0); + lean_ctor_release(x_94, 1); + x_96 = x_94; +} else { + lean_dec_ref(x_94); + x_96 = lean_box(0); } +x_97 = lean_box(0); +if (lean_is_scalar(x_96)) { + x_98 = lean_alloc_ctor(0, 2, 0); +} else { + x_98 = x_96; +} +lean_ctor_set(x_98, 0, x_97); +lean_ctor_set(x_98, 1, x_95); +return x_98; } } -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) { +} +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__16(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { if (lean_obj_tag(x_1) == 0) @@ -10612,21 +10772,21 @@ 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) { +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__11___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_11 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Command_runLinters___spec__15(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); +x_13 = l_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__16(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, uint8_t x_4, lean_object* 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_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__11___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, 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) { _start: { lean_object* x_14; uint8_t x_15; @@ -10639,7 +10799,7 @@ if (x_7 == 0) { lean_object* x_16; lean_object* x_17; x_16 = lean_box(0); -x_17 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12___lambda__1(x_1, x_2, x_3, x_4, x_5, x_10, x_16, x_11, x_12, x_13); +x_17 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__11___lambda__1(x_1, x_2, x_3, x_4, x_5, x_10, x_16, x_11, x_12, x_13); return x_17; } else @@ -10666,7 +10826,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 = lean_box(0); -x_29 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12___lambda__1(x_1, x_2, x_3, x_4, x_5, x_27, x_28, x_11, x_12, x_13); +x_29 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__11___lambda__1(x_1, x_2, x_3, x_4, x_5, x_27, x_28, x_11, x_12, x_13); return x_29; } } @@ -10694,12 +10854,12 @@ 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_box(0); -x_41 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12___lambda__1(x_1, x_2, x_3, x_4, x_5, x_39, x_40, x_11, x_12, x_13); +x_41 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__11___lambda__1(x_1, x_2, x_3, x_4, x_5, x_39, x_40, x_11, x_12, x_13); return x_41; } } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12___lambda__3(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, uint8_t x_6, double x_7, lean_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___at_Lean_Elab_Command_runLinters___spec__11___lambda__3(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, uint8_t x_6, double x_7, 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; @@ -10722,7 +10882,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_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12___lambda__2(x_1, x_2, x_15, x_3, x_4, x_5, x_6, x_7, lean_box(0), x_18, x_11, x_12, x_19); +x_20 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__11___lambda__2(x_1, x_2, x_15, x_3, x_4, x_5, x_6, x_7, lean_box(0), x_18, x_11, x_12, x_19); lean_dec(x_12); lean_dec(x_4); return x_20; @@ -10734,19 +10894,19 @@ x_21 = lean_ctor_get(x_17, 1); lean_inc(x_21); lean_dec(x_17); x_22 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__4___closed__2; -x_23 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12___lambda__2(x_1, x_2, x_15, x_3, x_4, x_5, x_6, x_7, lean_box(0), x_22, x_11, x_12, x_21); +x_23 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__11___lambda__2(x_1, x_2, x_15, x_3, x_4, x_5, x_6, x_7, lean_box(0), x_22, x_11, x_12, x_21); lean_dec(x_12); lean_dec(x_4); return x_23; } } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12___lambda__4(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, uint8_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__11___lambda__4(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, uint8_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_dec(x_8); -x_12 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__14___rarg(x_10, x_11); +x_12 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__13___rarg(x_10, x_11); x_13 = lean_ctor_get(x_12, 0); lean_inc(x_13); x_14 = lean_ctor_get(x_12, 1); @@ -10756,7 +10916,7 @@ x_15 = lean_alloc_closure((void*)(l_Lean_withTraceNode___at_Lean_Elab_Command_ru lean_closure_set(x_15, 0, x_1); lean_inc(x_10); lean_inc(x_9); -x_16 = l_Lean_withSeconds___at_Lean_Elab_Command_runLinters___spec__15(x_15, x_9, x_10, x_14); +x_16 = l_Lean_withSeconds___at_Lean_Elab_Command_runLinters___spec__14(x_15, x_9, x_10, x_14); if (lean_obj_tag(x_16) == 0) { lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; lean_object* x_68; uint8_t x_69; @@ -10816,7 +10976,7 @@ x_28 = lean_st_ref_set(x_10, x_23, x_24); x_29 = lean_ctor_get(x_28, 1); lean_inc(x_29); lean_dec(x_28); -x_30 = l_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__18(x_19, x_9, x_10, x_29); +x_30 = l_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__16(x_19, x_9, x_10, x_29); lean_dec(x_10); lean_dec(x_9); lean_dec(x_19); @@ -10902,7 +11062,7 @@ x_50 = lean_st_ref_set(x_10, x_49, x_24); x_51 = lean_ctor_get(x_50, 1); lean_inc(x_51); lean_dec(x_50); -x_52 = l_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__18(x_19, x_9, x_10, x_51); +x_52 = l_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__16(x_19, x_9, x_10, x_51); lean_dec(x_10); lean_dec(x_9); lean_dec(x_19); @@ -10962,7 +11122,7 @@ lean_object* x_61; double x_62; lean_object* x_63; x_61 = lean_box(0); x_62 = lean_unbox_float(x_20); lean_dec(x_20); -x_63 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12___lambda__3(x_13, x_2, x_3, x_19, x_4, x_21, x_62, x_5, lean_box(0), x_61, x_9, x_10, x_18); +x_63 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__11___lambda__3(x_13, x_2, x_3, x_19, x_4, x_21, x_62, x_5, lean_box(0), x_61, x_9, x_10, x_18); return x_63; } } @@ -10972,7 +11132,7 @@ lean_object* x_64; double x_65; lean_object* x_66; x_64 = lean_box(0); x_65 = lean_unbox_float(x_20); lean_dec(x_20); -x_66 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12___lambda__3(x_13, x_2, x_3, x_19, x_4, x_21, x_65, x_5, lean_box(0), x_64, x_9, x_10, x_18); +x_66 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__11___lambda__3(x_13, x_2, x_3, x_19, x_4, x_21, x_65, x_5, lean_box(0), x_64, x_9, x_10, x_18); return x_66; } } @@ -11007,7 +11167,7 @@ return x_77; } } } -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) { +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__11(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; @@ -11027,7 +11187,7 @@ 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__13(x_1, x_5, x_6, x_10); +x_15 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__12(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); @@ -11098,7 +11258,7 @@ 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__4(x_3, x_1, x_4, x_14, x_2, x_31, lean_box(0), x_30, x_5, x_6, x_18); +x_32 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__11___lambda__4(x_3, x_1, x_4, x_14, x_2, x_31, lean_box(0), x_30, x_5, x_6, x_18); return x_32; } } @@ -11111,12 +11271,12 @@ 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__4(x_3, x_1, x_4, x_14, x_2, x_35, lean_box(0), x_34, x_5, x_6, x_33); +x_36 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__11___lambda__4(x_3, x_1, x_4, x_14, x_2, x_35, lean_box(0), x_34, x_5, x_6, x_33); return x_36; } } } -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) { +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Command_runLinters___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) { _start: { lean_object* x_8; lean_object* x_9; @@ -11125,11 +11285,11 @@ 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__19(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Command_runLinters___spec__17(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_profileitM___at_Lean_Elab_Command_runLinters___spec__19___rarg___boxed), 7, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_profileitM___at_Lean_Elab_Command_runLinters___spec__17___rarg___boxed), 7, 0); return x_2; } } @@ -11183,7 +11343,7 @@ 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); +x_16 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__10(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) { @@ -11262,7 +11422,7 @@ 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); +x_33 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__10(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) { @@ -11370,14 +11530,14 @@ 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); +x_17 = lean_alloc_closure((void*)(l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__11___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); +x_20 = l_Lean_profileitM___at_Lean_Elab_Command_runLinters___spec__17___rarg(x_18, x_11, x_17, x_19, x_2, x_3, x_7); lean_dec(x_11); return x_20; } @@ -11433,18 +11593,6 @@ 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: { @@ -11456,11 +11604,11 @@ 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) { +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__9___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); +x_5 = l_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__9(x_1, x_2, x_3, x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); @@ -11532,18 +11680,18 @@ x_9 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4(x_1, x_2, 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) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__10___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _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); +x_6 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__10___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) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___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) { _start: { size_t x_10; size_t x_11; lean_object* x_12; @@ -11551,87 +11699,75 @@ 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); +x_12 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__10(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) { +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__13(x_1, x_2, x_3, x_4); +x_5 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__12(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) { +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__13___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); +x_3 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__13___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) { +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__13___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); +x_2 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__13(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) { +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Command_runLinters___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: { 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); +x_10 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Command_runLinters___spec__15(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) { +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___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_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__18(x_1, x_2, x_3, x_4); +x_5 = l_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__16(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) { +LEAN_EXPORT lean_object* l_Lean_withTraceNode___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, 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); +x_12 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__11___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, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__11___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) { _start: { uint8_t x_14; uint8_t x_15; double x_16; lean_object* x_17; @@ -11641,13 +11777,13 @@ x_15 = lean_unbox(x_7); lean_dec(x_7); x_16 = lean_unbox_float(x_8); lean_dec(x_8); -x_17 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12___lambda__2(x_1, x_2, x_3, x_14, x_5, x_6, x_15, x_16, x_9, x_10, x_11, x_12, x_13); +x_17 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__11___lambda__2(x_1, x_2, x_3, x_14, x_5, x_6, x_15, x_16, x_9, x_10, x_11, x_12, x_13); lean_dec(x_12); lean_dec(x_5); return x_17; } } -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, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__11___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; uint8_t x_15; double x_16; lean_object* x_17; @@ -11657,11 +11793,11 @@ x_15 = lean_unbox(x_6); lean_dec(x_6); x_16 = lean_unbox_float(x_7); lean_dec(x_7); -x_17 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12___lambda__3(x_1, x_2, x_14, x_4, x_5, x_15, x_16, x_8, x_9, x_10, x_11, x_12, x_13); +x_17 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__11___lambda__3(x_1, x_2, x_14, x_4, x_5, x_15, x_16, x_8, x_9, x_10, x_11, x_12, x_13); return x_17; } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12___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_withTraceNode___at_Lean_Elab_Command_runLinters___spec__11___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: { uint8_t x_12; uint8_t x_13; lean_object* x_14; @@ -11669,25 +11805,25 @@ x_12 = lean_unbox(x_3); lean_dec(x_3); x_13 = lean_unbox(x_6); lean_dec(x_6); -x_14 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12___lambda__4(x_1, x_2, x_12, x_4, x_5, x_13, x_7, x_8, x_9, x_10, x_11); +x_14 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__11___lambda__4(x_1, x_2, x_12, x_4, x_5, x_13, x_7, x_8, x_9, x_10, x_11); return x_14; } } -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) { +LEAN_EXPORT lean_object* l_Lean_withTraceNode___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) { _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); +x_9 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__11(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) { +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Command_runLinters___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) { _start: { 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); +x_8 = l_Lean_profileitM___at_Lean_Elab_Command_runLinters___spec__17___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; @@ -15706,7 +15842,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_runLinters___spec__13(x_9, x_2, x_3, x_4); +x_11 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__12(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); @@ -17328,7 +17464,7 @@ lean_closure_set(x_10, 1, x_5); x_11 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2971____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); +x_14 = lean_alloc_closure((void*)(l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__11___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); @@ -18091,7 +18227,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_runLinters___spec__13(x_21, x_6, x_7, x_8); +x_22 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__12(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); @@ -18448,7 +18584,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_runLinters___spec__13(x_21, x_6, x_7, x_8); +x_22 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__12(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); @@ -20670,7 +20806,7 @@ x_18 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommandTopLevel___lamb lean_closure_set(x_18, 0, x_1); x_19 = l_Lean_Elab_Command_elabCommandTopLevel___closed__1; x_20 = lean_box(0); -x_21 = l_Lean_profileitM___at_Lean_Elab_Command_runLinters___spec__19___rarg(x_19, x_17, x_18, x_20, x_2, x_3, x_13); +x_21 = l_Lean_profileitM___at_Lean_Elab_Command_runLinters___spec__17___rarg(x_19, x_17, x_18, x_20, x_2, x_3, x_13); lean_dec(x_17); return x_21; } @@ -20720,7 +20856,7 @@ x_37 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommandTopLevel___lamb lean_closure_set(x_37, 0, x_1); x_38 = l_Lean_Elab_Command_elabCommandTopLevel___closed__1; x_39 = lean_box(0); -x_40 = l_Lean_profileitM___at_Lean_Elab_Command_runLinters___spec__19___rarg(x_38, x_36, x_37, x_39, x_29, x_3, x_32); +x_40 = l_Lean_profileitM___at_Lean_Elab_Command_runLinters___spec__17___rarg(x_38, x_36, x_37, x_39, x_29, x_3, x_32); lean_dec(x_36); return x_40; } @@ -32031,10 +32167,10 @@ l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__4___c lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__4___closed__2); l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__5___closed__1 = _init_l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__5___closed__1(); lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__5___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_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__10___lambda__1___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__10___lambda__1___closed__1(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__10___lambda__1___closed__1); +l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__10___lambda__1___closed__2 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__10___lambda__1___closed__2(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__10___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(); diff --git a/stage0/stdlib/Lean/Elab/Deriving/DecEq.c b/stage0/stdlib/Lean/Elab/Deriving/DecEq.c index fd89c6b27ca2..7fac9322fb89 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/DecEq.c +++ b/stage0/stdlib/Lean/Elab/Deriving/DecEq.c @@ -35,6 +35,7 @@ static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__20; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkAuxFunctions___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*); 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_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__12(lean_object*, lean_object*, lean_object*, lean_object*); 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*); @@ -314,7 +315,6 @@ static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1 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; static lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4710____closed__6; -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*); @@ -9256,7 +9256,7 @@ x_215 = l_Lean_Syntax_node6(x_21, x_214, x_34, x_36, x_30, x_30, x_52, x_213); x_216 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__3; x_217 = l_Lean_Syntax_node2(x_21, x_216, x_32, x_215); x_218 = l_Lean_Elab_Deriving_DecEq_mkDecEqCmds___closed__5; -x_219 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__13(x_218, x_2, x_3, x_27); +x_219 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__12(x_218, x_2, x_3, x_27); x_220 = lean_ctor_get(x_219, 0); lean_inc(x_220); x_221 = lean_unbox(x_220); diff --git a/stage0/stdlib/Lean/Elab/LetRec.c b/stage0/stdlib/Lean/Elab/LetRec.c index 7852eba8fc37..a28b99b61bb5 100644 --- a/stage0/stdlib/Lean/Elab/LetRec.c +++ b/stage0/stdlib/Lean/Elab/LetRec.c @@ -13,85 +13,76 @@ #ifdef __cplusplus extern "C" { #endif +static lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetRec___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabLetRec___spec__2(size_t, size_t, lean_object*); lean_object* lean_string_utf8_extract(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__2___closed__3; +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__1___closed__3; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__23___closed__1; 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_object* lean_format_pretty(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__29___lambda__4___closed__8; +static lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___closed__3; +static lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__1___closed__3; static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__6___closed__20; lean_object* l_Lean_MapDeclarationExtension_insert___rarg(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__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_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__14___rarg___closed__1; static lean_object* l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__5___closed__1; -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__7; +LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__15___rarg(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetRec_declRange___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetRec___closed__3; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__14(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetRec___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_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__1(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__19___closed__5; extern lean_object* l_Lean_declRangeExt; lean_object* l_Lean_throwError___at_Lean_Elab_Term_mkAuxName___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_LetRec_0__Lean_Elab_Term_withAuxLocalDecls(lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__22___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_maxRecDepthErrorMessage; -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__1___closed__3; lean_object* l_Lean_indentD(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabLetRec(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___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_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__25(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_formatStxAux(lean_object*, uint8_t, lean_object*, lean_object*); lean_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_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___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_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___closed__2; lean_object* l_Lean_Name_toString(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_registerLetRecsToLift___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_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__9; +static lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__2___closed__1; lean_object* l_Lean_Syntax_getId(lean_object*); lean_object* l_Lean_Elab_WF_TerminationHints_rememberExtraParams(lean_object*, lean_object*, 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*); static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__6___closed__10; static lean_object* l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__9; static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__6___closed__2; -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__27___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___closed__1; lean_object* l_Lean_instBEqLocalInstance___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_elabAttrs___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__23___boxed(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__19___lambda__2___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetRec___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetRec___closed__1; lean_object* l_Lean_Syntax_getArgs(lean_object*); static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__6___closed__7; lean_object* l_Lean_replaceRef(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__25___boxed(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__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getPos_x3f(lean_object*, uint8_t); lean_object* lean_array_fget(lean_object*, lean_object*); static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__6___closed__18; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_elabLetRecDeclValues___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_getTailPos_x3f(lean_object*, uint8_t); -static lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_registerLetRecsToLift___spec__2___boxed(lean_object**); lean_object* l_Lean_Expr_fvarId_x21(lean_object*); -static lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___closed__1; 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*); -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__1; -static lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__1___closed__1; -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21(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*); lean_object* l_Lean_addBuiltinDeclarationRanges(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__24(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_forInUnsafe_loop___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_registerLetRecsToLift___spec__1___lambda__1(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__10; +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__24(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__19___lambda__2___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_isOfKind(lean_object*, lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_List_any___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -102,112 +93,105 @@ lean_object* l_Lean_Elab_Term_registerCustomErrorIfMVar(lean_object*, lean_objec static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__6___closed__17; lean_object* l_instBEqProd___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_string_utf8_byte_size(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__27___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_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*); lean_object* l_instHashableArray___rarg___boxed(lean_object*, lean_object*); static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__6___closed__6; -LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__24___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__6___closed__4; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabLetRec___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_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__2___closed__2; lean_object* l_List_appendTR___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___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_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__27(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_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___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_of_nat(lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwError___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*, lean_object*); static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__6___closed__19; static lean_object* l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__1; +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__8; static lean_object* l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__4___closed__2; -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__25___rarg(lean_object*); static lean_object* l_Lean_Elab_expandOptDocComment_x3f___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__1___closed__2; -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__1___closed__1; -static lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___closed__3; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabLetRec___spec__2___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_take(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_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__6___closed__5; lean_object* l_Lean_Elab_expandMacroImpl_x3f(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__28___lambda__4___closed__1; lean_object* l_Lean_ResolveName_resolveNamespace(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__20___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_unzip___rarg(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_registerLetRecsToLift___spec__1___lambda__1___boxed(lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofSyntax(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_elabAttrs___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*); static lean_object* l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__1___closed__4; +LEAN_EXPORT lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__2(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetRec_declRange___closed__2; +LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetRec_declRange___closed__4; -LEAN_EXPORT uint8_t l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__1(lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__3; lean_object* l_Lean_Elab_Term_expandOptType(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__15___rarg___closed__1; +LEAN_EXPORT uint8_t l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_toAttributeKind___boxed(lean_object*, lean_object*, 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*); static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__6___closed__8; lean_object* l_Lean_Syntax_getKind(lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_registerLetRecsToLift___spec__1___closed__4; +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__10; 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; -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__4; -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__22(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_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*); lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Term_expandDeclId___spec__6(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_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___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_EXPORT lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetRec_declRange___closed__6; 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_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__3(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* 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___regBuiltin_Lean_Elab_Term_elabLetRec_declRange___closed__7; lean_object* lean_array_to_list(lean_object*, lean_object*); extern lean_object* l_Lean_Elab_Term_termElabAttribute; +static lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___closed__2; 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*); lean_object* l_Lean_Name_append(lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_getDeclName_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__26___boxed(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_elabAttrs___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_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_instInhabitedExpr; -static lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___closed__6; 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_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__1___boxed(lean_object*); 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_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__1___closed__1; -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___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*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getSepArgs(lean_object*); -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_instHashableLocalInstance___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___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_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_instHashableProd___rarg___boxed(lean_object*, lean_object*, lean_object*); uint8_t lean_name_eq(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__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__11; static lean_object* l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__1___closed__2; -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__1___closed__2; +static lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__2___closed__4; static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__6___closed__15; static lean_object* l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__5___closed__2; -static lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__24___closed__2; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29(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_getArg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___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_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___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*); uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_mkFreshExprMVarImpl(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_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*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___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___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__1(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__19___closed__6; +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__1___closed__1; static lean_object* l_Lean_Elab_expandOptDocComment_x3f___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__1___closed__1; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_registerLetRecsToLift___spec__1___closed__2; extern lean_object* l_Std_Format_defWidth; @@ -216,118 +200,130 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_LetRec_0__Lean_Elab_Term_withAuxL 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; uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__15___rarg___closed__2; +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21(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__16___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__16(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_LetRec_0__Lean_Elab_Term_elabLetRecDeclValues___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_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_elabLetRecDeclValues___spec__1___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__3; -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__2; +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__5; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_registerLetRecsToLift___spec__1___closed__1; -LEAN_EXPORT lean_object* l_Lean_Elab_elabDeclAttrs___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*); +LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__1___boxed(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__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*); -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__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_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__22(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; lean_object* l_Lean_FileMap_leanPosToLspPos(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__2___boxed(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__24___boxed(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__28___lambda__4___closed__6; +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__24___rarg(lean_object*); lean_object* l_Lean_Elab_Term_addLocalVarInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__23(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__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___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_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___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___regBuiltin_Lean_Elab_Term_elabLetRec_declRange___closed__1; 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_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__2(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__28___lambda__4___closed__4; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); lean_object* lean_environment_main_module(lean_object*); static lean_object* l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__8; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___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* 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_Array_forInUnsafe_loop___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28(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_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__1___closed__2; 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*); -LEAN_EXPORT lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__6; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___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_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__2; lean_object* l_Lean_Meta_mkForallFVars(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___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*); +static lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__1___closed__2; 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*); 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_Lean_Syntax_isNone(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__1___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___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_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_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___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_Meta_InfoCacheKey_instHashableInfoCacheKey___boxed(lean_object*); -static lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__2___closed__4; +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__2; +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__26(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_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__23(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__25___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_LetRec_0__Lean_Elab_Term_withAuxLocalDecls_loop(lean_object*); static lean_object* l_Lean_Elab_expandOptDocComment_x3f___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__1___closed__3; -LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__26(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__19___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetRec___closed__5; lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__1___closed__3; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_erase_macro_scopes(lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__5; static lean_object* l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__4; lean_object* l_List_reverse___rarg(lean_object*); lean_object* l_Lean_Elab_Term_withAuxDecl___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_addDeclarationRanges___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__6___closed__16; +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__23___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__9; static lean_object* l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__1___closed__5; -static lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_addAuxDeclarationRanges___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__1___closed__7; -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetRec___closed__6; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_registerLetRecsToLift___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_array_uget(lean_object*, size_t); +LEAN_EXPORT lean_object* l_Lean_Elab_elabAttr___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_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___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_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__14___rarg(lean_object*); +static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__14___rarg___closed__2; 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*); -static lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__1___closed__3; 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*); +static lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__23___closed__2; lean_object* l_Lean_Name_mkStr4(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__14; static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__6___closed__11; static lean_object* l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__7; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___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* lean_string_append(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___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_object* l_Lean_getAttributeImpl(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___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_EXPORT lean_object* l_Lean_Elab_elabDeclAttrs___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*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); extern lean_object* l_Lean_Elab_unsupportedSyntaxExceptionId; -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkLambdaFVars(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, 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_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__2(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Exception_isRuntime(lean_object*); lean_object* l_Array_instBEqArray___rarg___boxed(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__28___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__26___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_expandMacros(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_elabAttrs___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*); lean_object* l_Lean_Expr_mvarId_x21(lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__25(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__1___closed__6; -static lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__24___closed__1; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); +static lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__2___closed__2; lean_object* l_Lean_MessageData_ofName(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_elabLetRecDeclValues___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*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabLetRec_declRange(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__2(lean_object*, lean_object*, lean_object*, 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_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__4___closed__1; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_registerLetRecsToLift___spec__1___closed__3; -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__10; 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*); +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__3; 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*); -static lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__2___closed__1; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___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_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*); +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__7; 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: @@ -1532,30 +1528,6 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__14(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_lt(x_4, x_3); -if (x_6 == 0) -{ -return x_5; -} -else -{ -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_7 = lean_array_uget(x_5, x_4); -x_8 = lean_unsigned_to_nat(0u); -x_9 = lean_array_uset(x_5, x_4, x_8); -x_10 = 1; -x_11 = lean_usize_add(x_4, x_10); -x_12 = lean_array_uset(x_9, x_4, x_7); -x_4 = x_11; -x_5 = x_12; -goto _start; -} -} -} static lean_object* _init_l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__1___closed__1() { _start: { @@ -1973,33 +1945,15 @@ static lean_object* _init_l_Lean_Elab_WF_elabTerminationHints___at___private_Lea _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("ident", 5); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__9() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__8; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__10() { -_start: -{ -lean_object* x_1; x_1 = lean_mk_string_from_bytes("no extra parameters bounds, please omit the `=>`", 48); return x_1; } } -static lean_object* _init_l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__11() { +static lean_object* _init_l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__9() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__10; +x_1 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__8; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } @@ -2135,54 +2089,47 @@ lean_inc(x_46); x_47 = l_Lean_Syntax_matchesNull(x_46, x_45); if (x_47 == 0) { -lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; size_t x_52; size_t x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; lean_object* x_57; lean_object* x_58; +lean_object* x_48; lean_object* x_49; uint8_t x_50; lean_object* x_51; lean_object* x_52; x_48 = l_Lean_Syntax_getArg(x_15, x_30); -x_49 = lean_box(0); -x_50 = l_Lean_Syntax_getArgs(x_46); +x_49 = l_Lean_Syntax_getArgs(x_46); lean_dec(x_46); -x_51 = lean_array_get_size(x_50); -x_52 = lean_usize_of_nat(x_51); -lean_dec(x_51); -x_53 = 0; -x_54 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__9; -x_55 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__14(x_49, x_54, x_52, x_53, x_50); -x_56 = 0; -x_57 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_57, 0, x_15); -lean_ctor_set(x_57, 1, x_55); -lean_ctor_set(x_57, 2, x_48); -lean_ctor_set_uint8(x_57, sizeof(void*)*3, x_56); -lean_ctor_set(x_3, 0, x_57); -x_58 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__1(x_1, x_4, x_2, x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_58; +x_50 = 0; +x_51 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_51, 0, x_15); +lean_ctor_set(x_51, 1, x_49); +lean_ctor_set(x_51, 2, x_48); +lean_ctor_set_uint8(x_51, sizeof(void*)*3, x_50); +lean_ctor_set(x_3, 0, x_51); +x_52 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__1(x_1, x_4, x_2, x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_52; } else { -lean_object* x_59; lean_object* x_60; uint8_t x_61; +lean_object* x_53; lean_object* x_54; uint8_t x_55; lean_dec(x_46); lean_free_object(x_3); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_59 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__11; -x_60 = l_Lean_throwErrorAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__12(x_15, x_59, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -x_61 = !lean_is_exclusive(x_60); -if (x_61 == 0) +x_53 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__9; +x_54 = l_Lean_throwErrorAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__12(x_15, x_53, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_55 = !lean_is_exclusive(x_54); +if (x_55 == 0) { -return x_60; +return x_54; } else { -lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_62 = lean_ctor_get(x_60, 0); -x_63 = lean_ctor_get(x_60, 1); -lean_inc(x_63); -lean_inc(x_62); -lean_dec(x_60); -x_64 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_64, 0, x_62); -lean_ctor_set(x_64, 1, x_63); -return x_64; +lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_54, 0); +x_57 = lean_ctor_get(x_54, 1); +lean_inc(x_57); +lean_inc(x_56); +lean_dec(x_54); +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; } } } @@ -2190,177 +2137,170 @@ return x_64; } else { -lean_object* x_65; lean_object* x_66; uint8_t x_67; -x_65 = lean_ctor_get(x_3, 0); -lean_inc(x_65); +lean_object* x_59; lean_object* x_60; uint8_t x_61; +x_59 = lean_ctor_get(x_3, 0); +lean_inc(x_59); lean_dec(x_3); -x_66 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__2; -lean_inc(x_65); -x_67 = l_Lean_Syntax_isOfKind(x_65, x_66); -if (x_67 == 0) +x_60 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__2; +lean_inc(x_59); +x_61 = l_Lean_Syntax_isOfKind(x_59, x_60); +if (x_61 == 0) { -lean_object* x_68; uint8_t x_69; -x_68 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__4; -lean_inc(x_65); -x_69 = l_Lean_Syntax_isOfKind(x_65, x_68); -if (x_69 == 0) +lean_object* x_62; uint8_t x_63; +x_62 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__4; +lean_inc(x_59); +x_63 = l_Lean_Syntax_isOfKind(x_59, x_62); +if (x_63 == 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_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_70 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__6; -x_71 = l_Lean_throwErrorAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__12(x_65, x_70, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -x_72 = lean_ctor_get(x_71, 0); -lean_inc(x_72); -x_73 = lean_ctor_get(x_71, 1); -lean_inc(x_73); -if (lean_is_exclusive(x_71)) { - lean_ctor_release(x_71, 0); - lean_ctor_release(x_71, 1); - x_74 = x_71; +x_64 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__6; +x_65 = l_Lean_throwErrorAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__12(x_59, x_64, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_66 = lean_ctor_get(x_65, 0); +lean_inc(x_66); +x_67 = lean_ctor_get(x_65, 1); +lean_inc(x_67); +if (lean_is_exclusive(x_65)) { + lean_ctor_release(x_65, 0); + lean_ctor_release(x_65, 1); + x_68 = x_65; } else { - lean_dec_ref(x_71); - x_74 = lean_box(0); + lean_dec_ref(x_65); + x_68 = lean_box(0); } -if (lean_is_scalar(x_74)) { - x_75 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_68)) { + x_69 = lean_alloc_ctor(1, 2, 0); } else { - x_75 = x_74; + x_69 = x_68; } -lean_ctor_set(x_75, 0, x_72); -lean_ctor_set(x_75, 1, x_73); -return x_75; +lean_ctor_set(x_69, 0, x_66); +lean_ctor_set(x_69, 1, x_67); +return x_69; } else { -lean_object* x_76; lean_object* x_77; -lean_dec(x_65); -x_76 = lean_box(0); -x_77 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__1(x_1, x_4, x_2, x_76, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_77; +lean_object* x_70; lean_object* x_71; +lean_dec(x_59); +x_70 = lean_box(0); +x_71 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__1(x_1, x_4, x_2, x_70, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_71; } } else { -lean_object* x_78; lean_object* x_79; lean_object* x_80; uint8_t x_81; -x_78 = lean_unsigned_to_nat(1u); -x_79 = l_Lean_Syntax_getArg(x_65, x_78); -x_80 = lean_unsigned_to_nat(2u); -lean_inc(x_79); -x_81 = l_Lean_Syntax_matchesNull(x_79, x_80); -if (x_81 == 0) +lean_object* x_72; lean_object* x_73; lean_object* x_74; uint8_t x_75; +x_72 = lean_unsigned_to_nat(1u); +x_73 = l_Lean_Syntax_getArg(x_59, x_72); +x_74 = lean_unsigned_to_nat(2u); +lean_inc(x_73); +x_75 = l_Lean_Syntax_matchesNull(x_73, x_74); +if (x_75 == 0) { -lean_object* x_82; uint8_t x_83; -x_82 = lean_unsigned_to_nat(0u); -x_83 = l_Lean_Syntax_matchesNull(x_79, x_82); -if (x_83 == 0) +lean_object* x_76; uint8_t x_77; +x_76 = lean_unsigned_to_nat(0u); +x_77 = l_Lean_Syntax_matchesNull(x_73, x_76); +if (x_77 == 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_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_84 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__6; -x_85 = l_Lean_throwErrorAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__12(x_65, x_84, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -x_86 = lean_ctor_get(x_85, 0); -lean_inc(x_86); -x_87 = lean_ctor_get(x_85, 1); -lean_inc(x_87); -if (lean_is_exclusive(x_85)) { - lean_ctor_release(x_85, 0); - lean_ctor_release(x_85, 1); - x_88 = x_85; +x_78 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__6; +x_79 = l_Lean_throwErrorAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__12(x_59, x_78, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_80 = lean_ctor_get(x_79, 0); +lean_inc(x_80); +x_81 = lean_ctor_get(x_79, 1); +lean_inc(x_81); +if (lean_is_exclusive(x_79)) { + lean_ctor_release(x_79, 0); + lean_ctor_release(x_79, 1); + x_82 = x_79; } else { - lean_dec_ref(x_85); - x_88 = lean_box(0); + lean_dec_ref(x_79); + x_82 = lean_box(0); } -if (lean_is_scalar(x_88)) { - x_89 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_82)) { + x_83 = 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); + x_83 = x_82; +} +lean_ctor_set(x_83, 0, x_80); +lean_ctor_set(x_83, 1, x_81); +return x_83; +} +else +{ +lean_object* x_84; lean_object* x_85; uint8_t x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_84 = l_Lean_Syntax_getArg(x_59, x_74); +x_85 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__7; +x_86 = 0; +x_87 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_87, 0, x_59); +lean_ctor_set(x_87, 1, x_85); +lean_ctor_set(x_87, 2, x_84); +lean_ctor_set_uint8(x_87, sizeof(void*)*3, x_86); +x_88 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_88, 0, x_87); +x_89 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__1(x_1, x_4, x_2, x_88, x_5, x_6, x_7, x_8, x_9, x_10, x_11); return x_89; } +} +else +{ +lean_object* x_90; lean_object* x_91; uint8_t x_92; +x_90 = lean_unsigned_to_nat(0u); +x_91 = l_Lean_Syntax_getArg(x_73, x_90); +lean_dec(x_73); +lean_inc(x_91); +x_92 = l_Lean_Syntax_matchesNull(x_91, x_90); +if (x_92 == 0) +{ +lean_object* x_93; lean_object* x_94; uint8_t x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_93 = l_Lean_Syntax_getArg(x_59, x_74); +x_94 = l_Lean_Syntax_getArgs(x_91); +lean_dec(x_91); +x_95 = 0; +x_96 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_96, 0, x_59); +lean_ctor_set(x_96, 1, x_94); +lean_ctor_set(x_96, 2, x_93); +lean_ctor_set_uint8(x_96, sizeof(void*)*3, x_95); +x_97 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_97, 0, x_96); +x_98 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__1(x_1, x_4, x_2, x_97, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_98; +} else { -lean_object* x_90; lean_object* x_91; uint8_t x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; -x_90 = l_Lean_Syntax_getArg(x_65, x_80); -x_91 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__7; -x_92 = 0; -x_93 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_93, 0, x_65); -lean_ctor_set(x_93, 1, x_91); -lean_ctor_set(x_93, 2, x_90); -lean_ctor_set_uint8(x_93, sizeof(void*)*3, x_92); -x_94 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_94, 0, x_93); -x_95 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__1(x_1, x_4, x_2, x_94, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_95; -} -} -else -{ -lean_object* x_96; lean_object* x_97; uint8_t x_98; -x_96 = lean_unsigned_to_nat(0u); -x_97 = l_Lean_Syntax_getArg(x_79, x_96); -lean_dec(x_79); -lean_inc(x_97); -x_98 = l_Lean_Syntax_matchesNull(x_97, x_96); -if (x_98 == 0) -{ -lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; size_t x_103; size_t x_104; lean_object* x_105; lean_object* x_106; uint8_t x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; -x_99 = l_Lean_Syntax_getArg(x_65, x_80); -x_100 = lean_box(0); -x_101 = l_Lean_Syntax_getArgs(x_97); -lean_dec(x_97); -x_102 = lean_array_get_size(x_101); -x_103 = lean_usize_of_nat(x_102); -lean_dec(x_102); -x_104 = 0; -x_105 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__9; -x_106 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__14(x_100, x_105, x_103, x_104, x_101); -x_107 = 0; -x_108 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_108, 0, x_65); -lean_ctor_set(x_108, 1, x_106); -lean_ctor_set(x_108, 2, x_99); -lean_ctor_set_uint8(x_108, sizeof(void*)*3, x_107); -x_109 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_109, 0, x_108); -x_110 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__1(x_1, x_4, x_2, x_109, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_110; -} -else -{ -lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; -lean_dec(x_97); +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_91); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_111 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__11; -x_112 = l_Lean_throwErrorAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__12(x_65, x_111, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -x_113 = lean_ctor_get(x_112, 0); -lean_inc(x_113); -x_114 = lean_ctor_get(x_112, 1); -lean_inc(x_114); -if (lean_is_exclusive(x_112)) { - lean_ctor_release(x_112, 0); - lean_ctor_release(x_112, 1); - x_115 = x_112; +x_99 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__9; +x_100 = l_Lean_throwErrorAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__12(x_59, x_99, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_101 = lean_ctor_get(x_100, 0); +lean_inc(x_101); +x_102 = lean_ctor_get(x_100, 1); +lean_inc(x_102); +if (lean_is_exclusive(x_100)) { + lean_ctor_release(x_100, 0); + lean_ctor_release(x_100, 1); + x_103 = x_100; } else { - lean_dec_ref(x_112); - x_115 = lean_box(0); + lean_dec_ref(x_100); + x_103 = lean_box(0); } -if (lean_is_scalar(x_115)) { - x_116 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_103)) { + x_104 = lean_alloc_ctor(1, 2, 0); } else { - x_116 = x_115; + x_104 = x_103; } -lean_ctor_set(x_116, 0, x_113); -lean_ctor_set(x_116, 1, x_114); -return x_116; +lean_ctor_set(x_104, 0, x_101); +lean_ctor_set(x_104, 1, x_102); +return x_104; } } } @@ -2687,7 +2627,7 @@ return x_14; } } } -static lean_object* _init_l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__15___rarg___closed__1() { +static lean_object* _init_l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__14___rarg___closed__1() { _start: { lean_object* x_1; @@ -2695,38 +2635,38 @@ x_1 = l_Lean_Elab_unsupportedSyntaxExceptionId; return x_1; } } -static lean_object* _init_l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__15___rarg___closed__2() { +static lean_object* _init_l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__14___rarg___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__15___rarg___closed__1; +x_2 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__14___rarg___closed__1; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__15___rarg(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__14___rarg(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__15___rarg___closed__2; +x_2 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__14___rarg___closed__2; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___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_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___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_7; -x_7 = lean_alloc_closure((void*)(l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__15___rarg), 1, 0); +x_7 = lean_alloc_closure((void*)(l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__14___rarg), 1, 0); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___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_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___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; 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; @@ -2773,7 +2713,7 @@ return x_22; } } } -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__16(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__15(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { uint8_t x_10; @@ -2786,7 +2726,7 @@ 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_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__17(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_13 = l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__16(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); @@ -2837,7 +2777,7 @@ lean_ctor_set(x_27, 8, x_22); lean_ctor_set(x_27, 9, x_23); lean_ctor_set(x_27, 10, x_24); lean_ctor_set_uint8(x_27, sizeof(void*)*11, x_25); -x_28 = l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__17(x_2, x_3, x_4, x_5, x_6, x_27, x_8, x_9); +x_28 = l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__16(x_2, x_3, x_4, x_5, x_6, x_27, x_8, x_9); lean_dec(x_8); lean_dec(x_27); lean_dec(x_6); @@ -2847,7 +2787,7 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__23(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___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) { _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; @@ -2894,7 +2834,7 @@ return x_22; } } } -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__22(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___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_object* x_9) { _start: { uint8_t x_10; @@ -2907,7 +2847,7 @@ 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_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__23(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_13 = l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__22(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); @@ -2958,7 +2898,7 @@ lean_ctor_set(x_27, 8, x_22); lean_ctor_set(x_27, 9, x_23); lean_ctor_set(x_27, 10, x_24); lean_ctor_set_uint8(x_27, sizeof(void*)*11, x_25); -x_28 = l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__23(x_2, x_3, x_4, x_5, x_6, x_27, x_8, x_9); +x_28 = l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__22(x_2, x_3, x_4, x_5, x_6, x_27, x_8, x_9); lean_dec(x_8); lean_dec(x_27); lean_dec(x_6); @@ -2968,7 +2908,7 @@ return x_28; } } } -static lean_object* _init_l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__24___closed__1() { +static lean_object* _init_l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__23___closed__1() { _start: { lean_object* x_1; lean_object* x_2; @@ -2978,21 +2918,21 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__24___closed__2() { +static lean_object* _init_l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__23___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__24___closed__1; +x_1 = l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__23___closed__1; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__24(lean_object* x_1, lean_object* x_2, lean_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_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__23(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_9 = l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__24___closed__2; +x_9 = l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__23___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); @@ -3002,26 +2942,26 @@ lean_ctor_set(x_11, 1, x_8); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__25___rarg(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__24___rarg(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__15___rarg___closed__2; +x_2 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__14___rarg___closed__2; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__25(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__24(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_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__25___rarg), 1, 0); +x_7 = lean_alloc_closure((void*)(l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__24___rarg), 1, 0); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__1(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___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___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; @@ -3190,7 +3130,7 @@ return x_41; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___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; @@ -3202,7 +3142,7 @@ lean_ctor_set(x_7, 1, x_4); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___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_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___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; @@ -3213,7 +3153,7 @@ lean_ctor_set(x_8, 1, x_6); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___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_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___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; @@ -3224,7 +3164,7 @@ lean_ctor_set(x_8, 1, x_6); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___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_Lean_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___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_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; @@ -3250,23 +3190,23 @@ 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_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__1___boxed), 4, 1); +x_19 = lean_alloc_closure((void*)(l_Lean_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___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_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__2___boxed), 4, 1); +x_21 = lean_alloc_closure((void*)(l_Lean_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___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_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___boxed), 6, 3); +x_22 = lean_alloc_closure((void*)(l_Lean_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___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_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___boxed), 6, 3); +x_23 = lean_alloc_closure((void*)(l_Lean_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___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); @@ -3439,14 +3379,14 @@ 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_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__22(x_67, x_72, x_2, x_3, x_4, x_5, x_6, x_7, x_27); +x_73 = l_Lean_throwErrorAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21(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_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__24(x_67, x_2, x_3, x_4, x_5, x_6, x_7, x_27); +x_74 = l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__23(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); @@ -3465,13 +3405,13 @@ 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_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__25___rarg(x_27); +x_75 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__24___rarg(x_27); return x_75; } } } } -LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___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_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__25(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _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; @@ -3518,7 +3458,7 @@ return x_22; } } } -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___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_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___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_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { uint8_t x_10; @@ -3582,7 +3522,7 @@ return x_28; } } } -static lean_object* _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__1___closed__1() { +static lean_object* _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__1___closed__1() { _start: { lean_object* x_1; @@ -3590,7 +3530,7 @@ x_1 = lean_mk_string_from_bytes("Term", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__1___closed__2() { +static lean_object* _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__1___closed__2() { _start: { lean_object* x_1; @@ -3598,23 +3538,23 @@ x_1 = lean_mk_string_from_bytes("byTactic", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__1___closed__3() { +static lean_object* _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__1___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__1___closed__1; x_2 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__1___closed__2; -x_3 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__1___closed__1; -x_4 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__1___closed__2; +x_3 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__1___closed__1; +x_4 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__1___closed__2; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -LEAN_EXPORT uint8_t l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__1(lean_object* x_1) { +LEAN_EXPORT uint8_t l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__1(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; -x_2 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__1___closed__3; +x_2 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__1___closed__3; x_3 = lean_name_eq(x_1, x_2); if (x_3 == 0) { @@ -3630,7 +3570,7 @@ return x_5; } } } -static lean_object* _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__2___closed__1() { +static lean_object* _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__2___closed__1() { _start: { lean_object* x_1; @@ -3638,16 +3578,16 @@ x_1 = lean_mk_string_from_bytes("unknown attribute [", 19); return x_1; } } -static lean_object* _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__2___closed__2() { +static lean_object* _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__2___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__2___closed__1; +x_1 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__2___closed__1; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__2___closed__3() { +static lean_object* _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__2___closed__3() { _start: { lean_object* x_1; @@ -3655,16 +3595,16 @@ x_1 = lean_mk_string_from_bytes("]", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__2___closed__4() { +static lean_object* _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__2___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__2___closed__3; +x_1 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__2___closed__3; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___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, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___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, lean_object* x_10) { _start: { lean_object* x_11; uint8_t x_12; @@ -3688,15 +3628,15 @@ lean_dec(x_16); lean_free_object(x_11); lean_dec(x_2); x_17 = l_Lean_MessageData_ofName(x_3); -x_18 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__2___closed__2; +x_18 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__2___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_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__2___closed__4; +x_20 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__2___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_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__26(x_21, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +x_22 = l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__25(x_21, x_4, x_5, x_6, x_7, x_8, x_9, x_14); lean_dec(x_8); lean_dec(x_6); return x_22; @@ -3736,15 +3676,15 @@ lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean lean_dec(x_27); lean_dec(x_2); x_28 = l_Lean_MessageData_ofName(x_3); -x_29 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__2___closed__2; +x_29 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__2___closed__2; x_30 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_30, 0, x_29); lean_ctor_set(x_30, 1, x_28); -x_31 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__2___closed__4; +x_31 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__2___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 = l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__26(x_32, x_4, x_5, x_6, x_7, x_8, x_9, x_25); +x_33 = l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__25(x_32, x_4, x_5, x_6, x_7, x_8, x_9, x_25); lean_dec(x_8); lean_dec(x_6); return x_33; @@ -3768,15 +3708,15 @@ return x_35; } } } -static lean_object* _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___closed__1() { +static lean_object* _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__1___boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__1___boxed), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___closed__2() { +static lean_object* _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___closed__2() { _start: { lean_object* x_1; @@ -3784,7 +3724,7 @@ x_1 = lean_mk_string_from_bytes("Attr", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___closed__3() { +static lean_object* _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___closed__3() { _start: { lean_object* x_1; @@ -3792,19 +3732,19 @@ x_1 = lean_mk_string_from_bytes("simple", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___closed__4() { +static lean_object* _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___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_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__1___closed__1; x_2 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__1___closed__2; -x_3 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___closed__2; -x_4 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___closed__3; +x_3 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___closed__2; +x_4 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___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_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___closed__5() { +static lean_object* _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___closed__5() { _start: { lean_object* x_1; @@ -3812,16 +3752,16 @@ x_1 = lean_mk_string_from_bytes("unknown attribute", 17); return x_1; } } -static lean_object* _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___closed__6() { +static lean_object* _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___closed__5; +x_1 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___closed__5; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___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_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___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_object* x_7, lean_object* x_8) { _start: { lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; @@ -3835,7 +3775,7 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_12 = l_Lean_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_12 = l_Lean_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8); if (lean_obj_tag(x_12) == 0) { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; @@ -3847,7 +3787,7 @@ lean_dec(x_12); x_15 = lean_unsigned_to_nat(1u); x_16 = l_Lean_Syntax_getArg(x_1, x_15); lean_dec(x_1); -x_17 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___closed__1; +x_17 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___closed__1; x_18 = lean_alloc_closure((void*)(l_Lean_expandMacros), 4, 2); lean_closure_set(x_18, 0, x_16); lean_closure_set(x_18, 1, x_17); @@ -3868,7 +3808,7 @@ lean_inc(x_21); lean_dec(x_19); lean_inc(x_20); x_22 = l_Lean_Syntax_getKind(x_20); -x_23 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___closed__4; +x_23 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___closed__4; x_24 = lean_name_eq(x_22, x_23); if (x_24 == 0) { @@ -3882,7 +3822,7 @@ x_26 = lean_box(0); x_27 = l_Lean_Name_str___override(x_26, x_25); x_28 = lean_unbox(x_13); lean_dec(x_13); -x_29 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__2(x_28, x_20, x_27, x_2, x_3, x_4, x_5, x_6, x_7, x_21); +x_29 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__2(x_28, x_20, x_27, x_2, x_3, x_4, x_5, x_6, x_7, x_21); lean_dec(x_7); lean_dec(x_5); lean_dec(x_3); @@ -3893,8 +3833,8 @@ else lean_object* x_30; lean_object* x_31; uint8_t x_32; lean_dec(x_22); lean_dec(x_13); -x_30 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___closed__6; -x_31 = l_Lean_throwErrorAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__27(x_20, x_30, x_2, x_3, x_4, x_5, x_6, x_7, x_21); +x_30 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___closed__6; +x_31 = l_Lean_throwErrorAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__26(x_20, x_30, x_2, x_3, x_4, x_5, x_6, x_7, x_21); lean_dec(x_7); lean_dec(x_5); lean_dec(x_4); @@ -3930,7 +3870,7 @@ lean_dec(x_36); x_38 = lean_erase_macro_scopes(x_37); x_39 = lean_unbox(x_13); lean_dec(x_13); -x_40 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__2(x_39, x_20, x_38, x_2, x_3, x_4, x_5, x_6, x_7, x_21); +x_40 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__2(x_39, x_20, x_38, x_2, x_3, x_4, x_5, x_6, x_7, x_21); lean_dec(x_7); lean_dec(x_5); lean_dec(x_3); @@ -3998,7 +3938,7 @@ return x_48; } } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__27(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { uint8_t x_12; @@ -4064,7 +4004,7 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_36 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20(x_14, x_5, x_6, x_7, x_8, x_35, x_10, x_11); +x_36 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19(x_14, x_5, x_6, x_7, x_8, x_35, x_10, x_11); if (lean_obj_tag(x_36) == 0) { lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; @@ -4364,7 +4304,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_elabAttrs___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___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_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Elab_elabAttrs___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___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; size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; @@ -4373,7 +4313,7 @@ x_10 = lean_usize_of_nat(x_9); lean_dec(x_9); x_11 = 0; x_12 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__7; -x_13 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28(x_1, x_10, x_11, x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_13 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__27(x_1, x_10, x_11, x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8); if (lean_obj_tag(x_13) == 0) { uint8_t x_14; @@ -4420,7 +4360,7 @@ return x_21; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_elabDeclAttrs___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___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_Elab_elabDeclAttrs___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___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) { _start: { lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; @@ -4429,12 +4369,12 @@ x_10 = l_Lean_Syntax_getArg(x_1, x_9); lean_dec(x_1); x_11 = l_Lean_Syntax_getSepArgs(x_10); lean_dec(x_10); -x_12 = l_Lean_Elab_elabAttrs___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_12 = l_Lean_Elab_elabAttrs___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__18(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_11); return x_12; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__1___closed__1() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__1___closed__1() { _start: { lean_object* x_1; @@ -4442,27 +4382,27 @@ x_1 = lean_mk_string_from_bytes("failed to infer 'let rec' declaration type", 42 return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__1___closed__2() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__1___closed__2() { _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__29___lambda__1___closed__1; +x_1 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___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_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__1___closed__3() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__1___closed__3() { _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__29___lambda__1___closed__2; +x_1 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__1___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_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___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_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___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; @@ -4482,7 +4422,7 @@ lean_inc(x_11); x_12 = lean_ctor_get(x_10, 1); lean_inc(x_12); lean_dec(x_10); -x_13 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__1___closed__3; +x_13 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__1___closed__3; x_14 = l_Lean_Elab_Term_registerCustomErrorIfMVar(x_11, x_1, x_13, x_3, x_4, x_5, x_6, x_7, x_8, x_12); lean_dec(x_4); lean_dec(x_3); @@ -4659,7 +4599,7 @@ return x_51; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { _start: { lean_object* x_17; lean_object* x_18; lean_object* x_19; @@ -4743,7 +4683,7 @@ return x_30; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___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, lean_object* x_8, lean_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_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___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, lean_object* x_8, 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; @@ -4833,7 +4773,7 @@ x_34 = l_Lean_Syntax_getArg(x_4, x_33); lean_inc(x_1); x_35 = l_Lean_Elab_Term_expandOptType(x_1, x_34); lean_dec(x_34); -x_36 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__1___boxed), 9, 1); +x_36 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__1___boxed), 9, 1); lean_closure_set(x_36, 0, x_35); lean_inc(x_13); lean_inc(x_12); @@ -4893,7 +4833,7 @@ lean_inc(x_54); x_55 = lean_ctor_get(x_53, 1); lean_inc(x_55); lean_dec(x_53); -x_56 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__2(x_5, x_1, x_2, x_15, x_20, x_41, x_40, x_46, x_54, x_8, x_9, x_10, x_11, x_12, x_13, x_55); +x_56 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__2(x_5, x_1, x_2, x_15, x_20, x_41, x_40, x_46, x_54, x_8, x_9, x_10, x_11, x_12, x_13, x_55); lean_dec(x_5); return x_56; } @@ -4945,7 +4885,7 @@ lean_dec(x_45); x_63 = lean_unsigned_to_nat(4u); x_64 = l_Lean_Syntax_getArg(x_4, x_63); lean_dec(x_4); -x_65 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__2(x_5, x_1, x_2, x_15, x_20, x_41, x_40, x_61, x_64, x_8, x_9, x_10, x_11, x_12, x_13, x_62); +x_65 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__2(x_5, x_1, x_2, x_15, x_20, x_41, x_40, x_61, x_64, x_8, x_9, x_10, x_11, x_12, x_13, x_62); lean_dec(x_5); return x_65; } @@ -5059,7 +4999,7 @@ return x_77; } } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__1() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__1() { _start: { lean_object* x_1; @@ -5067,19 +5007,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__29___lambda__4___closed__2() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___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_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__1___closed__1; x_2 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__1___closed__2; -x_3 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__1___closed__1; -x_4 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__1; +x_3 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__1___closed__1; +x_4 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___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__29___lambda__4___closed__3() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__3() { _start: { lean_object* x_1; @@ -5087,19 +5027,19 @@ 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__29___lambda__4___closed__4() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___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_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__1___closed__1; x_2 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__1___closed__2; -x_3 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__1___closed__1; -x_4 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__3; +x_3 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__1___closed__1; +x_4 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___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__29___lambda__4___closed__5() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__5() { _start: { lean_object* x_1; @@ -5107,16 +5047,16 @@ 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__29___lambda__4___closed__6() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___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__29___lambda__4___closed__5; +x_1 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___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__29___lambda__4___closed__7() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__7() { _start: { lean_object* x_1; @@ -5124,19 +5064,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__29___lambda__4___closed__8() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___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_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__1___closed__1; x_2 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__1___closed__2; -x_3 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__1___closed__1; -x_4 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__7; +x_3 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__1___closed__1; +x_4 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___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__29___lambda__4___closed__9() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__9() { _start: { lean_object* x_1; @@ -5144,16 +5084,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__29___lambda__4___closed__10() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___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__29___lambda__4___closed__9; +x_1 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___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__29___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_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___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; @@ -5162,19 +5102,19 @@ x_12 = l_Lean_Syntax_getArg(x_1, x_11); 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__29___lambda__4___closed__2; +x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___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__29___lambda__4___closed__4; +x_17 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___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_31; uint8_t x_32; -x_31 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__8; +x_31 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__8; lean_inc(x_14); x_32 = l_Lean_Syntax_isOfKind(x_14, x_31); if (x_32 == 0) @@ -5190,7 +5130,7 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_33 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__15___rarg(x_10); +x_33 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__14___rarg(x_10); return x_33; } else @@ -5221,7 +5161,7 @@ lean_dec(x_14); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_22 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__6; +x_22 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___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_7); @@ -5251,7 +5191,7 @@ else { 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__29___lambda__3(x_20, x_3, x_2, x_14, x_1, x_18, x_28, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_29 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__3(x_20, x_3, x_2, x_14, x_1, x_18, x_28, x_4, x_5, x_6, x_7, x_8, x_9, x_10); return x_29; } } @@ -5262,13 +5202,13 @@ lean_object* x_36; lean_object* x_37; lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_36 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__10; -x_37 = l_Lean_throwErrorAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__16(x_14, x_36, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_36 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__10; +x_37 = l_Lean_throwErrorAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__15(x_14, x_36, x_4, x_5, x_6, x_7, x_8, x_9, x_10); return x_37; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { uint8_t x_11; @@ -5323,7 +5263,7 @@ lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_24 = l_Lean_Elab_elabDeclAttrs___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__18(x_23, x_4, x_5, x_6, x_7, x_8, x_9, x_19); +x_24 = l_Lean_Elab_elabDeclAttrs___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__17(x_23, x_4, x_5, x_6, x_7, x_8, x_9, x_19); if (lean_obj_tag(x_24) == 0) { lean_object* x_25; lean_object* x_26; lean_object* x_27; @@ -5338,7 +5278,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__29___lambda__4(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__28___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; @@ -5428,7 +5368,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__29___lambda__4(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__28___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; @@ -5524,7 +5464,7 @@ x_14 = lean_array_get_size(x_13); x_15 = lean_usize_of_nat(x_14); lean_dec(x_14); x_16 = 0; -x_17 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29(x_15, x_16, x_13, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_17 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28(x_15, x_16, x_13, x_2, x_3, x_4, x_5, x_6, x_7, x_8); if (lean_obj_tag(x_17) == 0) { uint8_t x_18; @@ -5664,25 +5604,11 @@ lean_dec(x_3); return x_9; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__14___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_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__14(x_1, x_2, x_6, x_7, x_5); -lean_dec(x_2); -lean_dec(x_1); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___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_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__14___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { lean_object* x_7; -x_7 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__15(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__14(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); @@ -5692,11 +5618,11 @@ lean_dec(x_1); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___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_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___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: { lean_object* x_9; -x_9 = l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__17(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_9 = l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__16(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); @@ -5705,11 +5631,11 @@ lean_dec(x_3); return x_9; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__23___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___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) { _start: { lean_object* x_9; -x_9 = l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__23(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_9 = l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__22(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); @@ -5718,11 +5644,11 @@ lean_dec(x_3); return x_9; } } -LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__24___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__23___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_9; -x_9 = l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__24(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_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__23(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); @@ -5732,11 +5658,11 @@ lean_dec(x_2); return x_9; } } -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__25___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__24___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { lean_object* x_7; -x_7 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__25(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__24(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); @@ -5746,47 +5672,47 @@ lean_dec(x_1); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___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_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___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_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__1(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__1(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__2(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__2(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___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_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___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_7; -x_7 = l_Lean_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_Lean_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_5); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___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_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___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) { _start: { lean_object* x_7; -x_7 = l_Lean_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_Lean_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_5); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__26___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__25___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_9; -x_9 = l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__26(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_9 = l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__25(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); @@ -5795,11 +5721,11 @@ lean_dec(x_3); return x_9; } } -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__27___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__26___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; -x_10 = l_Lean_throwErrorAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__27(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_10 = l_Lean_throwErrorAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__26(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); @@ -5808,30 +5734,30 @@ lean_dec(x_1); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__1___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__1___boxed(lean_object* x_1) { _start: { uint8_t x_2; lean_object* x_3; -x_2 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__1(x_1); +x_2 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__1(x_1); lean_dec(x_1); x_3 = lean_box(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___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_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___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_1); lean_dec(x_1); -x_12 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__2(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_12 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__2(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_9); lean_dec(x_7); lean_dec(x_5); return x_12; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__27___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { size_t x_12; size_t x_13; lean_object* x_14; @@ -5839,49 +5765,49 @@ x_12 = lean_unbox_usize(x_2); lean_dec(x_2); x_13 = lean_unbox_usize(x_3); lean_dec(x_3); -x_14 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28(x_1, x_12, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_14 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__27(x_1, x_12, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_1); return x_14; } } -LEAN_EXPORT lean_object* l_Lean_Elab_elabAttrs___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___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_EXPORT lean_object* l_Lean_Elab_elabAttrs___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__18___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_9; -x_9 = l_Lean_Elab_elabAttrs___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_9 = l_Lean_Elab_elabAttrs___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__18(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_1); return x_9; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___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_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___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_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_10 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_2); return x_10; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___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_17; -x_17 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +x_17 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); lean_dec(x_1); return x_17; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___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_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___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; lean_object* x_16; x_15 = lean_unbox(x_6); lean_dec(x_6); -x_16 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__3(x_1, x_2, x_3, x_4, x_5, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +x_16 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__3(x_1, x_2, x_3, x_4, x_5, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); return x_16; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { size_t x_11; size_t x_12; lean_object* x_13; @@ -5889,7 +5815,7 @@ x_11 = lean_unbox_usize(x_1); lean_dec(x_1); x_12 = lean_unbox_usize(x_2); lean_dec(x_2); -x_13 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29(x_11, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_13 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28(x_11, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); return x_13; } } @@ -7747,7 +7673,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetRec___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_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__1___closed__1; x_2 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__1___closed__2; -x_3 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__1___closed__1; +x_3 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__1___closed__1; x_4 = l___regBuiltin_Lean_Elab_Term_elabLetRec___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -7775,7 +7701,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetRec___closed__5() 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_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__1___closed__1; x_2 = l___regBuiltin_Lean_Elab_Term_elabLetRec___closed__3; -x_3 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__1___closed__1; +x_3 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__1___closed__1; x_4 = l___regBuiltin_Lean_Elab_Term_elabLetRec___closed__4; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -8016,10 +7942,6 @@ l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab lean_mark_persistent(l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__8); l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__9 = _init_l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__9(); lean_mark_persistent(l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__9); -l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__10 = _init_l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__10(); -lean_mark_persistent(l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__10); -l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__11 = _init_l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__11(); -lean_mark_persistent(l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__11); l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__4___closed__1 = _init_l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__4___closed__1(); lean_mark_persistent(l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__4___closed__1); l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__4___closed__2 = _init_l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__4___closed__2(); @@ -8028,66 +7950,66 @@ l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab lean_mark_persistent(l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__5___closed__1); l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__5___closed__2 = _init_l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__5___closed__2(); lean_mark_persistent(l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__5___closed__2); -l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__15___rarg___closed__1 = _init_l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__15___rarg___closed__1(); -lean_mark_persistent(l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__15___rarg___closed__1); -l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__15___rarg___closed__2 = _init_l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__15___rarg___closed__2(); -lean_mark_persistent(l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__15___rarg___closed__2); -l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__24___closed__1 = _init_l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__24___closed__1(); -lean_mark_persistent(l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__24___closed__1); -l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__24___closed__2 = _init_l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__24___closed__2(); -lean_mark_persistent(l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__24___closed__2); -l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__1___closed__1 = _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__1___closed__1(); -lean_mark_persistent(l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__1___closed__1); -l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__1___closed__2 = _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__1___closed__2(); -lean_mark_persistent(l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__1___closed__2); -l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__1___closed__3 = _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__1___closed__3(); -lean_mark_persistent(l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__1___closed__3); -l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__2___closed__1 = _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__2___closed__1(); -lean_mark_persistent(l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__2___closed__1); -l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__2___closed__2 = _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__2___closed__2(); -lean_mark_persistent(l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__2___closed__2); -l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__2___closed__3 = _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__2___closed__3(); -lean_mark_persistent(l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__2___closed__3); -l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__2___closed__4 = _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__2___closed__4(); -lean_mark_persistent(l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___lambda__2___closed__4); -l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___closed__1 = _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___closed__1(); -lean_mark_persistent(l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___closed__1); -l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___closed__2 = _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___closed__2(); -lean_mark_persistent(l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___closed__2); -l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___closed__3 = _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___closed__3(); -lean_mark_persistent(l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___closed__3); -l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___closed__4 = _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___closed__4(); -lean_mark_persistent(l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___closed__4); -l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___closed__5 = _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___closed__5(); -lean_mark_persistent(l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___closed__5); -l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___closed__6 = _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___closed__6(); -lean_mark_persistent(l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__20___closed__6); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__1___closed__1 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__1___closed__1(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__1___closed__1); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__1___closed__2 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__1___closed__2(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__1___closed__2); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__1___closed__3 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__1___closed__3(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__1___closed__3); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__1 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__1(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__1); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__2 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__2(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__2); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__3 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__3(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__3); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__4 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__4(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__4); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__5 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__5(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__5); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__6 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__6(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__6); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__7 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__7(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__7); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__8 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__8(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__8); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__9 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__9(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__9); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__10 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__10(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__29___lambda__4___closed__10); +l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__14___rarg___closed__1 = _init_l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__14___rarg___closed__1(); +lean_mark_persistent(l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__14___rarg___closed__1); +l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__14___rarg___closed__2 = _init_l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__14___rarg___closed__2(); +lean_mark_persistent(l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__14___rarg___closed__2); +l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__23___closed__1 = _init_l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__23___closed__1(); +lean_mark_persistent(l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__23___closed__1); +l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__23___closed__2 = _init_l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__23___closed__2(); +lean_mark_persistent(l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__23___closed__2); +l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__1___closed__1 = _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__1___closed__1); +l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__1___closed__2 = _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__1___closed__2); +l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__1___closed__3 = _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__1___closed__3(); +lean_mark_persistent(l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__1___closed__3); +l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__2___closed__1 = _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__2___closed__1); +l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__2___closed__2 = _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__2___closed__2); +l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__2___closed__3 = _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__2___closed__3(); +lean_mark_persistent(l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__2___closed__3); +l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__2___closed__4 = _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__2___closed__4(); +lean_mark_persistent(l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___lambda__2___closed__4); +l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___closed__1 = _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___closed__1(); +lean_mark_persistent(l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___closed__1); +l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___closed__2 = _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___closed__2(); +lean_mark_persistent(l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___closed__2); +l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___closed__3 = _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___closed__3(); +lean_mark_persistent(l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___closed__3); +l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___closed__4 = _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___closed__4(); +lean_mark_persistent(l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___closed__4); +l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___closed__5 = _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___closed__5(); +lean_mark_persistent(l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___closed__5); +l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___closed__6 = _init_l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___closed__6(); +lean_mark_persistent(l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19___closed__6); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__1___closed__1 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__1___closed__1(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__1___closed__1); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__1___closed__2 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__1___closed__2(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__1___closed__2); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__1___closed__3 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__1___closed__3(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__1___closed__3); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__1 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__1(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__1); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__2 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__2(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__2); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__3 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__3(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__3); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__4 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__4(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__4); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__5 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__5(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__5); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__6 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__6(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__6); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__7 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__7(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__7); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__8 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__8(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__8); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__9 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__9(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__9); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__10 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___lambda__4___closed__10(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__28___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__1___closed__1 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_registerLetRecsToLift___spec__1___closed__1(); diff --git a/stage0/stdlib/Lean/Elab/MutualDef.c b/stage0/stdlib/Lean/Elab/MutualDef.c index 554f6ab92390..c602d3693d9b 100644 --- a/stage0/stdlib/Lean/Elab/MutualDef.c +++ b/stage0/stdlib/Lean/Elab/MutualDef.c @@ -417,7 +417,6 @@ LEAN_EXPORT lean_object* l_Array_sequenceMap___at___private_Lean_Elab_MutualDef_ LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at_Lean_Elab_Term_elabMutualDef___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_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*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__1___closed__1; 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_object* l_Lean_Elab_Command_getRef(lean_object*, lean_object*, lean_object*); @@ -650,7 +649,6 @@ static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_p 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_declValToTerminationHint___spec__8(lean_object*, lean_object*, size_t, size_t, lean_object*); 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*); @@ -8692,30 +8690,6 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__8(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_lt(x_4, x_3); -if (x_6 == 0) -{ -return x_5; -} -else -{ -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_7 = lean_array_uget(x_5, x_4); -x_8 = lean_unsigned_to_nat(0u); -x_9 = lean_array_uset(x_5, x_4, x_8); -x_10 = 1; -x_11 = lean_usize_add(x_4, x_10); -x_12 = lean_array_uset(x_9, x_4, x_7); -x_4 = x_11; -x_5 = x_12; -goto _start; -} -} -} static lean_object* _init_l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__1___closed__1() { _start: { @@ -9252,54 +9226,47 @@ lean_inc(x_46); x_47 = l_Lean_Syntax_matchesNull(x_46, x_45); if (x_47 == 0) { -lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; size_t x_52; size_t x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; lean_object* x_57; lean_object* x_58; +lean_object* x_48; lean_object* x_49; uint8_t x_50; lean_object* x_51; lean_object* x_52; x_48 = l_Lean_Syntax_getArg(x_15, x_30); -x_49 = lean_box(0); -x_50 = l_Lean_Syntax_getArgs(x_46); +x_49 = l_Lean_Syntax_getArgs(x_46); lean_dec(x_46); -x_51 = lean_array_get_size(x_50); -x_52 = lean_usize_of_nat(x_51); -lean_dec(x_51); -x_53 = 0; -x_54 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___closed__2; -x_55 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__8(x_49, x_54, x_52, x_53, x_50); -x_56 = 0; -x_57 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_57, 0, x_15); -lean_ctor_set(x_57, 1, x_55); -lean_ctor_set(x_57, 2, x_48); -lean_ctor_set_uint8(x_57, sizeof(void*)*3, x_56); -lean_ctor_set(x_3, 0, x_57); -x_58 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__1(x_1, x_4, x_2, x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_58; +x_50 = 0; +x_51 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_51, 0, x_15); +lean_ctor_set(x_51, 1, x_49); +lean_ctor_set(x_51, 2, x_48); +lean_ctor_set_uint8(x_51, sizeof(void*)*3, x_50); +lean_ctor_set(x_3, 0, x_51); +x_52 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__1(x_1, x_4, x_2, x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_52; } else { -lean_object* x_59; lean_object* x_60; uint8_t x_61; +lean_object* x_53; lean_object* x_54; uint8_t x_55; lean_dec(x_46); lean_free_object(x_3); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_59 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__2___closed__8; -x_60 = l_Lean_throwErrorAt___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__6(x_15, x_59, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -x_61 = !lean_is_exclusive(x_60); -if (x_61 == 0) +x_53 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__2___closed__8; +x_54 = l_Lean_throwErrorAt___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__6(x_15, x_53, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_55 = !lean_is_exclusive(x_54); +if (x_55 == 0) { -return x_60; +return x_54; } else { -lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_62 = lean_ctor_get(x_60, 0); -x_63 = lean_ctor_get(x_60, 1); -lean_inc(x_63); -lean_inc(x_62); -lean_dec(x_60); -x_64 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_64, 0, x_62); -lean_ctor_set(x_64, 1, x_63); -return x_64; +lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_54, 0); +x_57 = lean_ctor_get(x_54, 1); +lean_inc(x_57); +lean_inc(x_56); +lean_dec(x_54); +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; } } } @@ -9307,177 +9274,170 @@ return x_64; } else { -lean_object* x_65; lean_object* x_66; uint8_t x_67; -x_65 = lean_ctor_get(x_3, 0); -lean_inc(x_65); +lean_object* x_59; lean_object* x_60; uint8_t x_61; +x_59 = lean_ctor_get(x_3, 0); +lean_inc(x_59); lean_dec(x_3); -x_66 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__2___closed__2; -lean_inc(x_65); -x_67 = l_Lean_Syntax_isOfKind(x_65, x_66); -if (x_67 == 0) +x_60 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__2___closed__2; +lean_inc(x_59); +x_61 = l_Lean_Syntax_isOfKind(x_59, x_60); +if (x_61 == 0) { -lean_object* x_68; uint8_t x_69; -x_68 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__2___closed__4; -lean_inc(x_65); -x_69 = l_Lean_Syntax_isOfKind(x_65, x_68); -if (x_69 == 0) +lean_object* x_62; uint8_t x_63; +x_62 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__2___closed__4; +lean_inc(x_59); +x_63 = l_Lean_Syntax_isOfKind(x_59, x_62); +if (x_63 == 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_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_70 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__2___closed__6; -x_71 = l_Lean_throwErrorAt___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__6(x_65, x_70, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -x_72 = lean_ctor_get(x_71, 0); -lean_inc(x_72); -x_73 = lean_ctor_get(x_71, 1); -lean_inc(x_73); -if (lean_is_exclusive(x_71)) { - lean_ctor_release(x_71, 0); - lean_ctor_release(x_71, 1); - x_74 = x_71; +x_64 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__2___closed__6; +x_65 = l_Lean_throwErrorAt___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__6(x_59, x_64, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_66 = lean_ctor_get(x_65, 0); +lean_inc(x_66); +x_67 = lean_ctor_get(x_65, 1); +lean_inc(x_67); +if (lean_is_exclusive(x_65)) { + lean_ctor_release(x_65, 0); + lean_ctor_release(x_65, 1); + x_68 = x_65; } else { - lean_dec_ref(x_71); - x_74 = lean_box(0); + lean_dec_ref(x_65); + x_68 = lean_box(0); } -if (lean_is_scalar(x_74)) { - x_75 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_68)) { + x_69 = lean_alloc_ctor(1, 2, 0); } else { - x_75 = x_74; + x_69 = x_68; } -lean_ctor_set(x_75, 0, x_72); -lean_ctor_set(x_75, 1, x_73); -return x_75; +lean_ctor_set(x_69, 0, x_66); +lean_ctor_set(x_69, 1, x_67); +return x_69; } else { -lean_object* x_76; lean_object* x_77; -lean_dec(x_65); -x_76 = lean_box(0); -x_77 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__1(x_1, x_4, x_2, x_76, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_77; +lean_object* x_70; lean_object* x_71; +lean_dec(x_59); +x_70 = lean_box(0); +x_71 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__1(x_1, x_4, x_2, x_70, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_71; } } else { -lean_object* x_78; lean_object* x_79; lean_object* x_80; uint8_t x_81; -x_78 = lean_unsigned_to_nat(1u); -x_79 = l_Lean_Syntax_getArg(x_65, x_78); -x_80 = lean_unsigned_to_nat(2u); -lean_inc(x_79); -x_81 = l_Lean_Syntax_matchesNull(x_79, x_80); -if (x_81 == 0) +lean_object* x_72; lean_object* x_73; lean_object* x_74; uint8_t x_75; +x_72 = lean_unsigned_to_nat(1u); +x_73 = l_Lean_Syntax_getArg(x_59, x_72); +x_74 = lean_unsigned_to_nat(2u); +lean_inc(x_73); +x_75 = l_Lean_Syntax_matchesNull(x_73, x_74); +if (x_75 == 0) { -lean_object* x_82; uint8_t x_83; -x_82 = lean_unsigned_to_nat(0u); -x_83 = l_Lean_Syntax_matchesNull(x_79, x_82); -if (x_83 == 0) +lean_object* x_76; uint8_t x_77; +x_76 = lean_unsigned_to_nat(0u); +x_77 = l_Lean_Syntax_matchesNull(x_73, x_76); +if (x_77 == 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_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_84 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__2___closed__6; -x_85 = l_Lean_throwErrorAt___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__6(x_65, x_84, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -x_86 = lean_ctor_get(x_85, 0); -lean_inc(x_86); -x_87 = lean_ctor_get(x_85, 1); -lean_inc(x_87); -if (lean_is_exclusive(x_85)) { - lean_ctor_release(x_85, 0); - lean_ctor_release(x_85, 1); - x_88 = x_85; +x_78 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__2___closed__6; +x_79 = l_Lean_throwErrorAt___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__6(x_59, x_78, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_80 = lean_ctor_get(x_79, 0); +lean_inc(x_80); +x_81 = lean_ctor_get(x_79, 1); +lean_inc(x_81); +if (lean_is_exclusive(x_79)) { + lean_ctor_release(x_79, 0); + lean_ctor_release(x_79, 1); + x_82 = x_79; } else { - lean_dec_ref(x_85); - x_88 = lean_box(0); + lean_dec_ref(x_79); + x_82 = lean_box(0); } -if (lean_is_scalar(x_88)) { - x_89 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_82)) { + x_83 = lean_alloc_ctor(1, 2, 0); } else { - x_89 = x_88; + x_83 = x_82; } -lean_ctor_set(x_89, 0, x_86); -lean_ctor_set(x_89, 1, x_87); -return x_89; +lean_ctor_set(x_83, 0, x_80); +lean_ctor_set(x_83, 1, x_81); +return x_83; } else { -lean_object* x_90; lean_object* x_91; uint8_t x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; -x_90 = l_Lean_Syntax_getArg(x_65, x_80); -x_91 = l_Lean_Elab_instInhabitedDefViewElabHeader___closed__1; -x_92 = 0; -x_93 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_93, 0, x_65); -lean_ctor_set(x_93, 1, x_91); -lean_ctor_set(x_93, 2, x_90); -lean_ctor_set_uint8(x_93, sizeof(void*)*3, x_92); -x_94 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_94, 0, x_93); -x_95 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__1(x_1, x_4, x_2, x_94, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_95; +lean_object* x_84; lean_object* x_85; uint8_t x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_84 = l_Lean_Syntax_getArg(x_59, x_74); +x_85 = l_Lean_Elab_instInhabitedDefViewElabHeader___closed__1; +x_86 = 0; +x_87 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_87, 0, x_59); +lean_ctor_set(x_87, 1, x_85); +lean_ctor_set(x_87, 2, x_84); +lean_ctor_set_uint8(x_87, sizeof(void*)*3, x_86); +x_88 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_88, 0, x_87); +x_89 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__1(x_1, x_4, x_2, x_88, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_89; } } else { -lean_object* x_96; lean_object* x_97; uint8_t x_98; -x_96 = lean_unsigned_to_nat(0u); -x_97 = l_Lean_Syntax_getArg(x_79, x_96); -lean_dec(x_79); -lean_inc(x_97); -x_98 = l_Lean_Syntax_matchesNull(x_97, x_96); -if (x_98 == 0) +lean_object* x_90; lean_object* x_91; uint8_t x_92; +x_90 = lean_unsigned_to_nat(0u); +x_91 = l_Lean_Syntax_getArg(x_73, x_90); +lean_dec(x_73); +lean_inc(x_91); +x_92 = l_Lean_Syntax_matchesNull(x_91, x_90); +if (x_92 == 0) +{ +lean_object* x_93; lean_object* x_94; uint8_t x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_93 = l_Lean_Syntax_getArg(x_59, x_74); +x_94 = l_Lean_Syntax_getArgs(x_91); +lean_dec(x_91); +x_95 = 0; +x_96 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_96, 0, x_59); +lean_ctor_set(x_96, 1, x_94); +lean_ctor_set(x_96, 2, x_93); +lean_ctor_set_uint8(x_96, sizeof(void*)*3, x_95); +x_97 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_97, 0, x_96); +x_98 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__1(x_1, x_4, x_2, x_97, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_98; +} +else { -lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; size_t x_103; size_t x_104; lean_object* x_105; lean_object* x_106; uint8_t x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; -x_99 = l_Lean_Syntax_getArg(x_65, x_80); -x_100 = lean_box(0); -x_101 = l_Lean_Syntax_getArgs(x_97); -lean_dec(x_97); -x_102 = lean_array_get_size(x_101); -x_103 = lean_usize_of_nat(x_102); -lean_dec(x_102); -x_104 = 0; -x_105 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___closed__2; -x_106 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__8(x_100, x_105, x_103, x_104, x_101); -x_107 = 0; -x_108 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_108, 0, x_65); -lean_ctor_set(x_108, 1, x_106); -lean_ctor_set(x_108, 2, x_99); -lean_ctor_set_uint8(x_108, sizeof(void*)*3, x_107); -x_109 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_109, 0, x_108); -x_110 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__1(x_1, x_4, x_2, x_109, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_110; -} -else -{ -lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; -lean_dec(x_97); +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_91); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_111 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__2___closed__8; -x_112 = l_Lean_throwErrorAt___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__6(x_65, x_111, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -x_113 = lean_ctor_get(x_112, 0); -lean_inc(x_113); -x_114 = lean_ctor_get(x_112, 1); -lean_inc(x_114); -if (lean_is_exclusive(x_112)) { - lean_ctor_release(x_112, 0); - lean_ctor_release(x_112, 1); - x_115 = x_112; +x_99 = l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__2___closed__8; +x_100 = l_Lean_throwErrorAt___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__6(x_59, x_99, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_101 = lean_ctor_get(x_100, 0); +lean_inc(x_101); +x_102 = lean_ctor_get(x_100, 1); +lean_inc(x_102); +if (lean_is_exclusive(x_100)) { + lean_ctor_release(x_100, 0); + lean_ctor_release(x_100, 1); + x_103 = x_100; } else { - lean_dec_ref(x_112); - x_115 = lean_box(0); + lean_dec_ref(x_100); + x_103 = lean_box(0); } -if (lean_is_scalar(x_115)) { - x_116 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_103)) { + x_104 = lean_alloc_ctor(1, 2, 0); } else { - x_116 = x_115; + x_104 = x_103; } -lean_ctor_set(x_116, 0, x_113); -lean_ctor_set(x_116, 1, x_114); -return x_116; +lean_ctor_set(x_104, 0, x_101); +lean_ctor_set(x_104, 1, x_102); +return x_104; } } } @@ -9896,20 +9856,6 @@ lean_dec(x_3); return x_9; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__8___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_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__8(x_1, x_2, x_6, x_7, x_5); -lean_dec(x_2); -lean_dec(x_1); -return x_8; -} -} static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__1___closed__1() { _start: { diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/Main.c b/stage0/stdlib/Lean/Elab/PreDefinition/Main.c index dfc2edee317f..92cef423e0c8 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/Main.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/Main.c @@ -240,7 +240,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Util_SCC_0__Lean_SCC_updateLowLinkOf__ lean_object* l_Lean_Meta_mapErrorImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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*); static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_Main___hyg_1992____closed__11; -lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__13(size_t, size_t, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__15___lambda__2___closed__5; LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Elab_addPreDefinitions___spec__8(lean_object*, lean_object*); @@ -256,6 +255,7 @@ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_PreDefinit LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at___private_Lean_Elab_PreDefinition_Main_0__Lean_Elab_partitionPreDefs___spec__15(lean_object*, lean_object*); lean_object* l_Lean_indentExpr(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_addPreDefinitions___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_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__12(size_t, size_t, lean_object*); static lean_object* l_Lean_Core_transform_visit___at___private_Lean_Elab_PreDefinition_Main_0__Lean_Elab_betaReduceLetRecApps___spec__4___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_addPreDefinitions___spec__4___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_PreDefinition_Main_0__Lean_Elab_partitionPreDefs___spec__26(lean_object*, size_t, size_t, lean_object*); @@ -11485,7 +11485,7 @@ x_43 = lean_array_get_size(x_18); x_44 = lean_usize_of_nat(x_43); lean_dec(x_43); lean_inc(x_18); -x_45 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__13(x_44, x_4, x_18); +x_45 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__12(x_44, x_4, x_18); x_46 = lean_array_to_list(lean_box(0), x_45); x_47 = lean_box(0); x_48 = l_List_mapTR_loop___at_Lean_Meta_substCore___spec__6(x_46, x_47); diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/WF/GuessLex.c b/stage0/stdlib/Lean/Elab/PreDefinition/WF/GuessLex.c index b77d48e0219f..248d75c2fc36 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/WF/GuessLex.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/WF/GuessLex.c @@ -16,7 +16,6 @@ extern "C" { static lean_object* l_Lean_Loop_forIn_loop___at_Lean_Elab_WF_GuessLex_collectRecCalls___spec__5___closed__3; LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__3___rarg(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_Elab_WF_GuessLex_withRecApps_loop___spec__7___rarg(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_WF_guessLex___spec__14(lean_object*, 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_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_naryVarNames___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_loop___at_Lean_Elab_WF_GuessLex_evalRecCall___spec__3___lambda__9___closed__2; @@ -27,8 +26,9 @@ static lean_object* l_panic___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__ lean_object* l___private_Init_Util_0__outOfBounds___rarg(lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___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_lambdaTelescope___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__19___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__1___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___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_WF_GuessLex_RecCallWithContext_posString(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve_go___at_Lean_Elab_WF_guessLex___spec__9___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_originalVarNames(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_unresolveNameGlobal___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_GuessLex_0__Lean_Elab_WF_GuessLex_reprGuessLexRel____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_2280____closed__18; @@ -41,11 +41,12 @@ LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_RecCallCache_eval___boxed(lean_ static lean_object* l_Lean_Elab_WF_GuessLex_initFn____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_9____closed__4; LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__11(lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_GuessLex_GuessLexRel_toNatRel___closed__9; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getNatParams___spec__1___closed__1; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__4___closed__1; LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__14(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__1___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_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__4(size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__4(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_Elab_WF_GuessLex_explainMutualFailure___spec__5___closed__9; LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getForbiddenByTrivialSizeOf___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_naryVarNames___spec__1___closed__2; @@ -53,7 +54,7 @@ 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_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_guessLex___lambda__1___closed__4; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__8___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_guessLex___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_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_naryVarNames___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_loop___at_Lean_Elab_WF_GuessLex_evalRecCall___spec__3___lambda__9___closed__4; static lean_object* l_Lean_Elab_WF_guessLex___closed__1; @@ -74,23 +75,25 @@ LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_collectRecCalls___lambda__1(lea LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__28(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5(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_Elab_WF_GuessLex_evalRecCall___lambda__2___closed__5; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__13___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_name_append_after(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__28___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_Elab_WF_GuessLex_instToFormatGuessLexRel___closed__1; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__4___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___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_Elab_WF_GuessLex_collectRecCalls___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_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_buildTermWF___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_Array_foldlMUnsafe_fold___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__21___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_GuessLex_explainNonMutualFailure___closed__1; -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_buildTermWF___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_buildTermWF___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_x27_loop___at_Lean_Elab_WF_GuessLex_naryVarNames___spec__1___boxed(lean_object**); +lean_object* l_Lean_mkAppB(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__21___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___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_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform___spec__1(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_collectRecCalls___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__15___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_WF_GuessLex_mkTupleSyntax___closed__6; LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__8___rarg(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_List_forIn_loop___at_Lean_Elab_WF_GuessLex_evalRecCall___spec__3___lambda__9___closed__5; lean_object* l_Lean_FileMap_toPosition(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__2___boxed(lean_object*, lean_object*); @@ -117,6 +120,7 @@ lean_object* l_Lean_MessageData_ofList(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_WF_GuessLex_withRecApps_processApp___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__14___boxed(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*); static lean_object* l_List_forIn_loop___at_Lean_Elab_WF_GuessLex_evalRecCall___spec__3___closed__2; static lean_object* l_Lean_Elab_WF_GuessLex_generateMeasures___closed__1; @@ -140,6 +144,7 @@ extern uint8_t l_instInhabitedBool; static lean_object* l_Lean_Elab_WF_GuessLex_GuessLexRel_toNatRel___closed__29; static lean_object* l_Lean_Elab_WF_GuessLex_GuessLexRel_toNatRel___closed__25; LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_SavedLocalContext_run___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve_go___at_Lean_Elab_WF_guessLex___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_WF_GuessLex_initFn____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_9_(lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__21___rarg___lambda__2___closed__4; static lean_object* l_Lean_Elab_WF_GuessLex_explainFailure___closed__3; @@ -148,14 +153,15 @@ static lean_object* l_Lean_Elab_WF_GuessLex_initFn____x40_Lean_Elab_PreDefinitio static lean_object* l_Lean_Elab_WF_GuessLex_instToFormatGuessLexRel___closed__3; static lean_object* l_Lean_Elab_WF_GuessLex_instToStringGuessLexRel___closed__3; LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__7___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_WF_guessLex___spec__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__22___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___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_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_buildTermWF___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_Array_foldlMUnsafe_fold___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__21___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*); lean_object* l_Lean_replaceRef(lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_GuessLex_initFn____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_9____closed__8; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__22___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_WF_GuessLex_generateCombinations_x3f_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___closed__3; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_filterSubsumed___spec__6___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_List_forIn_loop___at_Lean_Elab_WF_GuessLex_evalRecCall___spec__3___lambda__9___closed__10; lean_object* l_Lean_Syntax_getPos_x3f(lean_object*, uint8_t); @@ -165,14 +171,17 @@ lean_object* lean_array_fget(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_GuessLex_0__Lean_Elab_WF_GuessLex_reprGuessLexRel____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_2280____closed__19; LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__33(lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_loop___at_Lean_Elab_WF_GuessLex_evalRecCall___spec__3___lambda__3___closed__1; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getSizeOfParams___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_array_fset(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getTailPos_x3f(lean_object*, uint8_t); lean_object* l_Lean_Meta_Match_MatcherInfo_arity(lean_object*); static lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__1___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_withoutModifyingState___at_Lean_Elab_WF_GuessLex_SavedLocalContext_run___spec__1(lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__2(lean_object*, lean_object*, size_t, size_t, 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_Meta_withLetDecl___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__33___rarg(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__27___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_WF_GuessLex_getNatParams___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_RecCallWithContext_create(lean_object*, 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_WF_GuessLex_withRecApps_loop___spec__32(lean_object*, lean_object*, lean_object*); lean_object* lean_environment_find(lean_object*, lean_object*); @@ -189,7 +198,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUnif lean_object* l_Lean_Meta_Tactic_TryThis_addSuggestion(lean_object*, 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_WF_GuessLex_withRecApps_loop___spec__31___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_PreDefinition_WF_GuessLex_0__Lean_Elab_WF_GuessLex_reprGuessLexRel____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_2280____closed__13; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__11___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_getConstInfo___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__3___rarg___closed__3; static lean_object* l_List_forIn_loop___at_Lean_Elab_WF_GuessLex_evalRecCall___spec__3___lambda__8___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__8___boxed(lean_object*, lean_object*, lean_object*); @@ -202,8 +210,8 @@ static lean_object* l_Lean_Elab_WF_GuessLex_withRecApps___rarg___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_SavedLocalContext_create(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__4___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_GuessLex_withRecApps___rarg___closed__6; -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve___at_Lean_Elab_WF_guessLex___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___closed__1; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_reportWF___spec__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___lambda__1___closed__3; lean_object* l_Lean_stringToMessageData(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_SavedLocalContext_run___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_RecCallCache_eval(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -212,12 +220,14 @@ static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainM LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_filterSubsumed___spec__5(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_withRecApps___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___private_Lean_Elab_PreDefinition_WF_GuessLex_0__Lean_Elab_WF_GuessLex_reprGuessLexRel____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_2280____closed__22; +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve_go___at_Lean_Elab_WF_guessLex___spec__9___lambda__1(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_WF_GuessLex_evalRecCall___spec__3___lambda__9___closed__8; static lean_object* l___private_Lean_Elab_PreDefinition_WF_GuessLex_0__Lean_Elab_WF_GuessLex_reprGuessLexRel____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_2280____closed__6; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___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_List_forIn_loop___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__16(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__6___rarg(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_WF_GuessLex_withRecApps_loop___spec__32___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*); +static lean_object* l_Lean_Elab_WF_GuessLex_getNatParams___closed__3; LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__11___boxed(lean_object*, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__21___rarg___lambda__2___closed__3; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -225,20 +235,21 @@ static lean_object* l_Lean_Elab_WF_GuessLex_evalRecCall___lambda__2___closed__4; LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_WF_GuessLex_evalRecCall___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_filterSubsumed___spec__4___rarg___lambda__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_GuessLex_GuessLexRel_toNatRel___closed__23; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___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_WF_GuessLex_GuessLexRel_toNatRel___closed__21; static lean_object* l_Lean_Elab_WF_GuessLex_collectRecCalls___lambda__2___closed__6; 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_Lean_Elab_Tactic_run(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_WF_GuessLex_withRecApps_loop___spec__25(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__8(size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getNatParams___spec__1(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_Elab_WF_GuessLex_instToFormatGuessLexRel___closed__2; LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_WF_GuessLex_evalRecCall___spec__3___lambda__2___boxed(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_WF_GuessLex_evalRecCall___spec__3___lambda__2___closed__1; static lean_object* l_Lean_Loop_forIn_loop___at_Lean_Elab_WF_GuessLex_collectRecCalls___spec__3___closed__4; +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_reportWF___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Loop_forIn_loop___at_Lean_Elab_WF_GuessLex_collectRecCalls___spec__3___closed__7; LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Elab_WF_GuessLex_withRecApps_processApp___spec__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_explainFailure(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_WF_GuessLex_getSizeOfParams___spec__2___lambda__1___closed__1; extern lean_object* l_instInhabitedNat; LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_WF_GuessLex_evalRecCall___spec__3___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_getForbiddenByTrivialSizeOf(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -249,6 +260,7 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_build LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_filterSubsumed___spec__4___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static size_t l_Lean_Elab_WF_GuessLex_generateMeasures___closed__4; static lean_object* l_Lean_Elab_WF_GuessLex_GuessLexRel_toNatRel___closed__19; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__10(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_filterPairsM___at_Lean_Elab_WF_GuessLex_filterSubsumed___spec__3___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_collectRecCalls___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_WF_guessLex___closed__2; @@ -265,10 +277,11 @@ static lean_object* l_Lean_Elab_WF_GuessLex_mkSizeOf___closed__3; static lean_object* l_Lean_Elab_WF_GuessLex_GuessLexRel_toNatRel___closed__20; LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_withRecApps_containsRecFn(lean_object*); static lean_object* l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__7; +lean_object* l_ReaderT_instApplicativeReaderT___rarg(lean_object*); static lean_object* l_Lean_Elab_WF_GuessLex_collectRecCalls___lambda__2___closed__3; static lean_object* l_Lean_Loop_forIn_loop___at_Lean_Elab_WF_GuessLex_collectRecCalls___spec__3___closed__2; size_t lean_usize_of_nat(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at_Lean_Elab_WF_GuessLex_buildTermWF___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_EXPORT lean_object* l_Array_mapIdxM_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_generateMeasures___spec__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_GuessLex_GuessLexRel_toNatRel___closed__16; LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_filterSubsumed___lambda__2(lean_object*, lean_object*); @@ -278,11 +291,12 @@ lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_lambdaTelescopeImp___rarg( LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_RecCallWithContext_posString___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_GuessLex_GuessLexRel_toNatRel___closed__6; static lean_object* l_Lean_Elab_WF_guessLex___lambda__1___closed__2; -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___lambda__1___closed__1; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___lambda__1___closed__1; static lean_object* l_Lean_Elab_WF_GuessLex_originalVarNames___closed__1; static lean_object* l_Lean_Elab_WF_GuessLex_GuessLexRel_toNatRel___closed__32; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__1(size_t, size_t, lean_object*); +uint8_t l_Array_contains___at_Lean_Meta_setMVarUserNamesAt___spec__1(lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); lean_object* l_Lean_Expr_getRevArg_x21(lean_object*, lean_object*); 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*); @@ -292,31 +306,32 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_gener LEAN_EXPORT lean_object* l_Lean_unresolveNameGlobal___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__2___boxed(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_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___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_Elab_WF_guessLex___spec__11(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_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__10___rarg(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_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_WF_GuessLex_evalRecCall___spec__3___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* l_Lean_Meta_Match_Extension_getMatcherInfo_x3f(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___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_Lean_Elab_WF_GuessLex_explainFailure___closed__1; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15(lean_object*, size_t, size_t, 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_Lean_getConstInfo___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__5(lean_object*); static lean_object* l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__3; static lean_object* l___private_Lean_Elab_PreDefinition_WF_GuessLex_0__Lean_Elab_WF_GuessLex_reprGuessLexRel____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_2280____closed__21; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___lambda__1___closed__4; lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); static lean_object* l_List_forIn_loop___at_Lean_Elab_WF_GuessLex_evalRecCall___spec__3___lambda__9___closed__1; lean_object* lean_nat_to_int(lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___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_MessageData_ofSyntax(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_WF_GuessLex_generateMeasures___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_GuessLex_collectRecCalls___lambda__3___closed__1; static lean_object* l_Lean_Elab_WF_GuessLex_filterSubsumed___lambda__2___closed__4; static lean_object* l___private_Lean_Elab_PreDefinition_WF_GuessLex_0__Lean_Elab_WF_GuessLex_reprGuessLexRel____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_2280____closed__7; -LEAN_EXPORT lean_object* l_Lean_Elab_WF_guessLex___lambda__2(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_WF_guessLex___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__9(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_go___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_WF_GuessLex_evalRecCall___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_Array_mapIdxM_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_withoutErrToSorryImp___rarg(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_WF_GuessLex_evalRecCall___spec__3___lambda__9___closed__9; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_WF_GuessLex_generateCombinations_x3f_go___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -325,6 +340,7 @@ LEAN_EXPORT uint8_t l_List_forIn_loop___at_Lean_Elab_WF_GuessLex_evalRecCall___s LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__28___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__30___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve_go___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_getSizeOfParams___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_GuessLex_GuessLexRel_toNatRel___closed__4; lean_object* l_Lean_getRevAliases(lean_object*, lean_object*); lean_object* l_Lean_Meta_SavedState_restore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -336,6 +352,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_WF_Guess lean_object* l_Lean_Name_getPrefix(lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainNonMutualFailure___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_WF_GuessLex_naryVarNames_freshen___closed__1; +LEAN_EXPORT lean_object* l_Lean_Elab_WF_guessLex___lambda__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*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__25___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__18___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_GuessLex_collectRecCalls___lambda__2___closed__4; @@ -346,28 +363,31 @@ LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_filterSubsumed(lean_object*); static lean_object* l_Lean_Elab_WF_GuessLex_filterSubsumed___lambda__2___closed__3; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_WF_GuessLex_withRecApps_processApp___spec__1(lean_object*); lean_object* l_Lean_Elab_goalsToMessageData(lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__2___lambda__1___closed__2; 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_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Elab_WF_GuessLex_filterSubsumed___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__9___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_collectRecCalls___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_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__3(lean_object*, 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_Elab_WF_guessLex___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__3(size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__5___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_guessLex___lambda__1___closed__5; lean_object* l_Lean_Meta_check(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Match_MatcherInfo_numAlts(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_instToFormatGuessLexRel(uint8_t); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7___closed__2; static lean_object* l___private_Lean_Elab_PreDefinition_WF_GuessLex_0__Lean_Elab_WF_GuessLex_reprGuessLexRel____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_2280____closed__20; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__24___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_WF_guessLex___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_RecCallCache_mk(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__2(lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_GuessLex_GuessLexRel_toNatRel___closed__30; uint8_t l_Lean_Expr_isLit(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_Elab_WF_GuessLex_solve_go___at_Lean_Elab_WF_guessLex___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getForbiddenByTrivialSizeOf___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_instMonadLiftReaderT(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__2; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__15(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__29___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__5; static lean_object* l_Lean_Elab_WF_GuessLex_GuessLexRel_toNatRel___closed__1; @@ -389,10 +409,13 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Elab_WF_GuessLex_with lean_object* lean_st_mk_ref(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__12___rarg(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_Elab_WF_GuessLex_solve_go___at_Lean_Elab_WF_guessLex___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_WF_GuessLex_withRecApps_loop___spec__4___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_getNatParams(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__9___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve_go___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_WF_GuessLex_getNatParams___closed__1; static lean_object* l_Lean_Elab_WF_GuessLex_GuessLexRel_toNatRel___closed__33; lean_object* l_Lean_Syntax_node3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_synthInstance(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -411,8 +434,8 @@ LEAN_EXPORT lean_object* l_Lean_unresolveNameGlobal_unresolveNameCore___at_Lean_ LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__11___rarg(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_GuessLex_0__Lean_Elab_WF_GuessLex_reprGuessLexRel____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_2280____closed__24; extern lean_object* l_Lean_levelZero; -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_guessLex___spec__12___boxed(lean_object**); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__7(lean_object*, size_t, size_t, 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_Elab_WF_GuessLex_reportWF___spec__2___lambda__1___closed__2; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__7(size_t, size_t, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__6(lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_trimTermWF___boxed(lean_object*, lean_object*); extern lean_object* l_Lean_instInhabitedExpr; @@ -425,11 +448,10 @@ LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_WF_GuessLex_evalRecCal static lean_object* l_Lean_Elab_WF_GuessLex_GuessLexRel_toNatRel___closed__28; LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_filterSubsumed___boxed(lean_object*); static lean_object* l_Lean_Elab_WF_GuessLex_initFn____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_9____closed__1; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___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_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___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*); static lean_object* l_Lean_Elab_WF_GuessLex_withRecApps___rarg___closed__5; static lean_object* l___private_Lean_Elab_PreDefinition_WF_GuessLex_0__Lean_Elab_WF_GuessLex_reprGuessLexRel____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_2280____closed__8; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___closed__3; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__5___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__10___boxed(lean_object*, lean_object*); lean_object* l_List_mapTR_loop___at_Lean_Meta_substCore___spec__6(lean_object*, lean_object*); @@ -443,6 +465,8 @@ static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_ static lean_object* l_List_forIn_loop___at_Lean_Elab_WF_GuessLex_evalRecCall___spec__3___lambda__8___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_withRecApps_containsRecFn___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_List_mapTR_loop___at_Lean_Elab_WF_guessLex___spec__16(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_guessLex___spec__11___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_reportWF___spec__1(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__5___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_componentsRev(lean_object*); LEAN_EXPORT lean_object* l_Lean_withoutModifyingState___at_Lean_Elab_WF_GuessLex_SavedLocalContext_run___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -453,6 +477,7 @@ LEAN_EXPORT lean_object* l_Lean_withoutModifyingState___at_Lean_Elab_WF_GuessLex lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); extern lean_object* l_Lean_rootNamespace; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___closed__6; +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve___at_Lean_Elab_WF_guessLex___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve_go(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_withRecApps___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* l_Lean_Meta_lambdaTelescope___at_Lean_PrettyPrinter_Delaborator_returnsPi___spec__1___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -468,7 +493,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_withRecApps_processRec___rarg(l static lean_object* l_Lean_Elab_WF_GuessLex_collectRecCalls___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7___closed__5; -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve_go___at_Lean_Elab_WF_guessLex___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_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__1___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*); static lean_object* l_Lean_Elab_WF_GuessLex_collectRecCalls___lambda__2___closed__2; @@ -482,13 +506,13 @@ LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at_Lean_Elab_WF_GuessLex_ static lean_object* l_Lean_Elab_WF_GuessLex_withRecApps___rarg___closed__7; static lean_object* l_List_forIn_loop___at_Lean_Elab_WF_GuessLex_evalRecCall___spec__3___lambda__9___closed__6; lean_object* l_Lean_Meta_saveState___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__6(size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__6(lean_object*, size_t, size_t, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at_Lean_Elab_WF_GuessLex_collectRecCalls___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_isEqvAux___at_Lean_Elab_WF_GuessLex_filterSubsumed___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_GuessLex_instToStringGuessLexRel___closed__1; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2___rarg___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withoutErrToSorry___at_Lean_Elab_WF_GuessLex_evalRecCall___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve_go___at_Lean_Elab_WF_guessLex___spec__10___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_Elab_WF_GuessLex_reportWF___spec__2___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_withRecApps_containsRecFn___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__17___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_GuessLex_0__Lean_Elab_WF_GuessLex_reprGuessLexRel____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_2280____closed__4; @@ -498,15 +522,15 @@ lean_object* lean_mk_syntax_ident(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__28___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___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_Lean_Elab_WF_guessLex___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_buildTermWF(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_buildTermWF(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_GuessLex_initFn____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_9____closed__5; static lean_object* l_Lean_Elab_WF_GuessLex_evalRecCall___lambda__3___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_RecCallCache_mk___boxed(lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Meta_instMonadMetaM; LEAN_EXPORT lean_object* l_Lean_Elab_WF_unpackMutualArg___at_Lean_Elab_WF_GuessLex_collectRecCalls___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_guessLex___spec__11___boxed(lean_object**); lean_object* l_Array_eraseIdx___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_unpackMutualArg___at_Lean_Elab_WF_GuessLex_collectRecCalls___spec__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_Elab_WF_guessLex___spec__15___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__8(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__19___rarg(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_Loop_forIn_loop___at_Lean_Elab_WF_GuessLex_collectRecCalls___spec__3___closed__1; @@ -529,37 +553,41 @@ lean_object* l_Array_append___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_GuessLex_mkSizeOf___closed__2; LEAN_EXPORT uint8_t l_Lean_Elab_WF_GuessLex_GuessLexRel_ofNat(lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___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_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__31___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* l_Lean_MessageData_ofExpr(lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_GuessLex_0__Lean_Elab_WF_GuessLex_reprGuessLexRel____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_2280____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainNonMutualFailure___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___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_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__7___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_guessLex___spec__12___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7___closed__3; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_filterSubsumed___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* l_Lean_throwError___at_Lean_Meta_mkSimpCongrTheorem___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__6___closed__1; static lean_object* l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__8; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__6___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___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_Loop_forIn_loop___at_Lean_Elab_WF_GuessLex_collectRecCalls___spec__3___closed__6; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__14(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_GuessLex_filterSubsumed___closed__1; static lean_object* l_Lean_Elab_WF_GuessLex_collectRecCalls___lambda__3___closed__2; lean_object* l_Array_extract___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__3(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_GuessLexRel_noConfusion(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_evalRecCall(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_hasAnyFVar_visit___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_guessLex___lambda__1___closed__7; LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve_go___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_loop___at_Lean_Elab_WF_GuessLex_evalRecCall___spec__3___closed__4; static lean_object* l_Lean_Elab_WF_GuessLex_generateMeasures___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_collectRecCalls___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_Lean_Elab_WF_GuessLex_generateMeasures___spec__2(size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7(lean_object*, lean_object*, lean_object*, size_t, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__13___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_mkSizeOf(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__2___lambda__1___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_Lean_Elab_WF_GuessLex_GuessLexRel_toCtorIdx(uint8_t); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__15(lean_object*); @@ -568,16 +596,17 @@ static lean_object* l_Lean_Elab_WF_GuessLex_withRecApps_processRec___rarg___clos LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Name_hasMacroScopes(lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___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_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getNatParams___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_List_forIn_loop___at_Lean_Elab_WF_GuessLex_evalRecCall___spec__3___lambda__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_guessLex(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_GuessLex_GuessLexRel_noConfusion___rarg___closed__1; +lean_object* l_instMonadControlT__1___rarg(lean_object*); static lean_object* l_Lean_Elab_WF_GuessLex_instReprGuessLexRel___closed__1; static lean_object* l_List_forIn_loop___at_Lean_Elab_WF_GuessLex_evalRecCall___spec__3___lambda__9___closed__11; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_explainNonMutualFailure___spec__1(size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_WF_GuessLex_GuessLexRel_toNatRel___closed__10; LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_GuessLex_0__Lean_Elab_WF_GuessLex_reprGuessLexRel____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_2280____closed__16; -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___closed__2; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__6___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_GuessLex_0__Lean_Elab_WF_GuessLex_reprGuessLexRel____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_2280____closed__12; static lean_object* l_panic___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__7___rarg___closed__3; @@ -586,11 +615,12 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_n LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_instToFormatGuessLexRel___boxed(lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__5___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_MatcherApp_transform___spec__6(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); lean_object* lean_string_length(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve_go___at_Lean_Elab_WF_guessLex___spec__10___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_WF_GuessLex_withRecApps___rarg___closed__1; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___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_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__3___boxed(lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___closed__1; static lean_object* l_Lean_Elab_WF_GuessLex_GuessLexRel_toNatRel___closed__2; static lean_object* l_Lean_Loop_forIn_loop___at_Lean_Elab_WF_GuessLex_collectRecCalls___spec__3___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -600,7 +630,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_explainFailure___boxed(lean_obj LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_SepArray_ofElems(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_filterSubsumed___spec__4___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_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__13(size_t, size_t, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7___closed__1; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_GuessLex_collectRecCalls___lambda__2___closed__5; @@ -608,8 +637,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_WF_guessLex___lambda__2___boxed(lean_object LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at_Lean_Elab_WF_GuessLex_trimTermWF___spec__1(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_WF_GuessLex_withRecApps_loop___spec__21___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* 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_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__1(size_t, size_t, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___lambda__1___closed__3; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__1___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7___closed__4; static lean_object* l_Lean_Elab_WF_GuessLex_GuessLexRel_toNatRel___closed__26; @@ -625,6 +653,7 @@ static lean_object* l_Lean_Loop_forIn_loop___at_Lean_Elab_WF_GuessLex_collectRec LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_filterSubsumed___spec__4(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_buildTermWF___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_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__12(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___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_Elab_WF_GuessLex_GuessLexRel_toNatRel___closed__24; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___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*); @@ -692,6 +721,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_WF_Guess LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_mkTupleSyntax(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__14___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_ofSubarray___rarg(lean_object*); +extern lean_object* l_Lean_Core_instMonadCoreM; static lean_object* l_Lean_Elab_WF_GuessLex_explainFailure___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_explainMutualFailure(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__1___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*); @@ -709,12 +739,12 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_f LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__4___rarg___boxed(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_throwError___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__4(lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___lambda__1___closed__4; LEAN_EXPORT lean_object* l_Lean_unresolveNameGlobal_unresolveNameCore___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_GuessLex_explainFailure___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_filterSubsumed___lambda__1___boxed(lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___closed__4; -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_WF_guessLex___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__12___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_inspectCall(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at_Lean_Elab_WF_GuessLex_collectRecCalls___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkHashMapImp___rarg(lean_object*); @@ -728,13 +758,14 @@ static lean_object* l_Lean_Elab_WF_GuessLex_GuessLexRel_toNatRel___closed__13; static lean_object* l_Lean_Elab_WF_GuessLex_initFn____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_9____closed__7; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__26___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_GuessLexRel_noConfusion___rarg___lambda__1(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve___at_Lean_Elab_WF_guessLex___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_explainMutualFailure___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_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__5___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_GuessLex_GuessLexRel_toNatRel___closed__14; LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_formatTable___boxed(lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__17(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_getNatParams___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Expr_hasFVar(lean_object*); lean_object* l_Lean_Elab_ensureNoRecFn(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_withRecApps_loop(lean_object*); extern lean_object* l_Lean_instInhabitedName; @@ -746,6 +777,7 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_sol LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__12___boxed(lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); LEAN_EXPORT lean_object* l_Lean_Elab_WF_unpackUnaryArg___at_Lean_Elab_WF_GuessLex_collectRecCalls___spec__4(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_Array_foldlMUnsafe_fold___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__30(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__6___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_GuessLex_0__Lean_Elab_WF_GuessLex_reprGuessLexRel____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_2280____closed__26; @@ -753,13 +785,17 @@ LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_generateMeasures___boxed(lean_o lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_WF_TerminationBy_unexpand(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_Elab_WF_GuessLex_solve___at_Lean_Elab_WF_guessLex___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_WF_guessLex___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* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_mkTupleSyntax___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_loop___at_Lean_Elab_WF_GuessLex_evalRecCall___spec__3___closed__3; lean_object* l_List_redLength___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve_go___rarg___lambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_guessLex___lambda__1___closed__6; +static lean_object* l_Lean_Elab_WF_GuessLex_getNatParams___closed__2; lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_Expr_hasAnyFVar_visit___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_GuessLex_evalRecCall___closed__1; lean_object* l_Lean_Elab_Tactic_evalTactic(lean_object*, 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_GuessLex_originalVarNames___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -768,7 +804,9 @@ LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_instReprGuessLexRel; LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_explainNonMutualFailure(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_string_append(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_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___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_Elab_WF_guessLex___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_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___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_WF_GuessLex_reportWF(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_filterSubsumed___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* l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_GuessLexRel_toNatRel(uint8_t); @@ -781,7 +819,6 @@ lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_resolveGlobalName___at_Lean_Elab_WF_GuessLex_naryVarNames_freshen___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_evalRecCall___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__13(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_guessLex___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*); static lean_object* l_Lean_Elab_WF_GuessLex_evalRecCall___lambda__3___closed__1; LEAN_EXPORT uint8_t l_Lean_Elab_WF_GuessLex_instDecidableEqGuessLexRel(uint8_t, uint8_t); static lean_object* l___private_Lean_Elab_PreDefinition_WF_GuessLex_0__Lean_Elab_WF_GuessLex_reprGuessLexRel____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_2280____closed__2; @@ -817,8 +854,10 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_build static lean_object* l_Lean_Elab_WF_GuessLex_RecCallWithContext_posString___closed__2; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__22(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_unresolveNameGlobal_unresolveNameCore___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__2___lambda__1___closed__3; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_WF_GuessLex_evalRecCall___spec__3___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_getSizeOfParams(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_mvarId_x21(lean_object*); lean_object* l_Lean_InductiveVal_numCtors(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1(lean_object*, lean_object*); @@ -837,11 +876,11 @@ LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at_Lean_Elab_WF_GuessLex_ lean_object* lean_expr_instantiate1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__20(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_inspectCall___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_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___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_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__5(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__1___rarg___lambda__1(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_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___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_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___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_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_guessLex___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__31(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Elab_WF_GuessLex_filterSubsumed___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__30___rarg(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*); @@ -852,8 +891,8 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_expla lean_object* l_Lean_Meta_etaExpand(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__4___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_unpackArg___at_Lean_Elab_WF_GuessLex_collectRecCalls___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_guessLex___spec__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_WF_GuessLex_GuessLexRel_toNatRel___closed__3; +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_getSizeOfParams___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7___closed__6; static lean_object* l_Lean_Elab_WF_GuessLex_collectRecCalls___lambda__2___closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_evalRecCall___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -17272,6248 +17311,9246 @@ lean_dec(x_3); return x_9; } } -LEAN_EXPORT uint8_t l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_isForbidden(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT uint8_t l_Lean_Expr_hasAnyFVar_visit___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__1(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_4; uint8_t x_5; -x_4 = lean_array_get_size(x_1); -x_5 = lean_nat_dec_lt(x_2, x_4); -lean_dec(x_4); -if (x_5 == 0) +uint8_t x_3; +x_3 = l_Lean_Expr_hasFVar(x_2); +if (x_3 == 0) { -uint8_t x_6; -x_6 = 0; -return x_6; +uint8_t x_4; +lean_dec(x_2); +x_4 = 0; +return x_4; } else { -lean_object* x_7; uint8_t x_8; -x_7 = lean_array_fget(x_1, x_2); -x_8 = l_Array_contains___at___private_Lean_Meta_FunInfo_0__Lean_Meta_collectDeps_visit___spec__2(x_7, x_3); -lean_dec(x_7); -return x_8; -} -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_isForbidden___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +switch (lean_obj_tag(x_2)) { +case 1: { -uint8_t x_4; lean_object* x_5; -x_4 = l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_isForbidden(x_1, x_2, x_3); -lean_dec(x_3); +lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_5 = lean_ctor_get(x_2, 0); +lean_inc(x_5); lean_dec(x_2); -lean_dec(x_1); -x_5 = lean_box(x_4); -return x_5; -} +x_6 = l_Lean_Expr_fvar___override(x_5); +x_7 = l_Array_contains___at_Lean_Meta_setMVarUserNamesAt___spec__1(x_1, x_6); +lean_dec(x_6); +return x_7; } -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform___spec__1(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { -_start: -{ -uint8_t x_5; -x_5 = lean_usize_dec_eq(x_3, x_4); -if (x_5 == 0) +case 5: { -lean_object* x_6; uint8_t x_7; -x_6 = lean_array_uget(x_2, x_3); -x_7 = l_Array_contains___at___private_Lean_Meta_FunInfo_0__Lean_Meta_collectDeps_visit___spec__2(x_6, x_1); -lean_dec(x_6); -if (x_7 == 0) +lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_8 = lean_ctor_get(x_2, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_2, 1); +lean_inc(x_9); +lean_dec(x_2); +x_10 = l_Lean_Expr_hasAnyFVar_visit___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__1(x_1, x_8); +if (x_10 == 0) { -size_t x_8; size_t x_9; -x_8 = 1; -x_9 = lean_usize_add(x_3, x_8); -x_3 = x_9; +x_2 = x_9; goto _start; } else { -uint8_t x_11; -x_11 = 1; -return x_11; -} -} -else -{ uint8_t x_12; -x_12 = 0; +lean_dec(x_9); +x_12 = 1; return x_12; } } -} -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform___spec__2(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { -_start: -{ -uint8_t x_5; -x_5 = lean_usize_dec_eq(x_3, x_4); -if (x_5 == 0) -{ -lean_object* x_6; uint8_t x_7; -x_6 = lean_array_uget(x_2, x_3); -x_7 = lean_nat_dec_lt(x_1, x_6); -lean_dec(x_6); -if (x_7 == 0) +case 6: { -uint8_t x_8; -x_8 = 1; -return x_8; -} -else +lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_13 = lean_ctor_get(x_2, 1); +lean_inc(x_13); +x_14 = lean_ctor_get(x_2, 2); +lean_inc(x_14); +lean_dec(x_2); +x_15 = l_Lean_Expr_hasAnyFVar_visit___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__1(x_1, x_13); +if (x_15 == 0) { -size_t x_9; size_t x_10; -x_9 = 1; -x_10 = lean_usize_add(x_3, x_9); -x_3 = x_10; +x_2 = x_14; goto _start; } -} else { -uint8_t x_12; -x_12 = 0; -return x_12; -} +uint8_t x_17; +lean_dec(x_14); +x_17 = 1; +return x_17; } } -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +case 7: { -lean_object* x_6; lean_object* x_7; lean_object* x_8; -lean_dec(x_4); -x_6 = lean_unsigned_to_nat(1u); -x_7 = lean_nat_add(x_1, x_6); -lean_dec(x_1); -x_8 = l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform(x_2, x_3, x_7, x_5); -return x_8; -} +lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_18 = lean_ctor_get(x_2, 1); +lean_inc(x_18); +x_19 = lean_ctor_get(x_2, 2); +lean_inc(x_19); +lean_dec(x_2); +x_20 = l_Lean_Expr_hasAnyFVar_visit___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__1(x_1, x_18); +if (x_20 == 0) +{ +x_2 = x_19; +goto _start; } -static lean_object* _init_l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform___closed__1() { -_start: +else { -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; +uint8_t x_22; +lean_dec(x_19); +x_22 = 1; +return x_22; } } -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: +case 8: { -lean_object* x_5; lean_object* x_6; lean_object* x_24; uint8_t x_25; -x_5 = lean_array_get_size(x_2); -x_24 = lean_unsigned_to_nat(0u); -x_25 = lean_nat_dec_lt(x_24, x_5); -if (x_25 == 0) +lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_23 = lean_ctor_get(x_2, 1); +lean_inc(x_23); +x_24 = lean_ctor_get(x_2, 2); +lean_inc(x_24); +x_25 = lean_ctor_get(x_2, 3); +lean_inc(x_25); +lean_dec(x_2); +x_26 = l_Lean_Expr_hasAnyFVar_visit___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__1(x_1, x_23); +if (x_26 == 0) { -lean_object* x_26; -x_26 = lean_box(0); -x_6 = x_26; -goto block_23; +uint8_t x_27; +x_27 = l_Lean_Expr_hasAnyFVar_visit___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__1(x_1, x_24); +if (x_27 == 0) +{ +x_2 = x_25; +goto _start; } else { -size_t x_27; size_t x_28; uint8_t x_29; -x_27 = 0; -x_28 = lean_usize_of_nat(x_5); -x_29 = l_Array_anyMUnsafe_any___at_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform___spec__2(x_3, x_2, x_27, x_28); -if (x_29 == 0) -{ -lean_object* x_30; -x_30 = lean_box(0); -x_6 = x_30; -goto block_23; +uint8_t x_29; +lean_dec(x_25); +x_29 = 1; +return x_29; +} } else { -lean_object* x_31; lean_object* x_32; -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_31 = l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform___closed__1; -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_4); -return x_32; +uint8_t x_30; +lean_dec(x_25); +lean_dec(x_24); +x_30 = 1; +return x_30; } } -block_23: -{ -lean_object* x_7; lean_object* x_8; uint8_t x_9; -lean_dec(x_6); -x_7 = lean_array_get_size(x_1); -x_8 = lean_unsigned_to_nat(0u); -x_9 = lean_nat_dec_lt(x_8, x_7); -if (x_9 == 0) +case 10: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -lean_dec(x_7); -lean_inc(x_3); -x_10 = lean_mk_array(x_5, x_3); -x_11 = lean_array_push(x_4, x_10); -x_12 = lean_box(0); -x_13 = l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform___lambda__1(x_3, x_1, x_2, x_12, x_11); -return x_13; +lean_object* x_31; +x_31 = lean_ctor_get(x_2, 1); +lean_inc(x_31); +lean_dec(x_2); +x_2 = x_31; +goto _start; } -else -{ -size_t x_14; size_t x_15; uint8_t x_16; -x_14 = 0; -x_15 = lean_usize_of_nat(x_7); -lean_dec(x_7); -x_16 = l_Array_anyMUnsafe_any___at_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform___spec__1(x_3, x_1, x_14, x_15); -if (x_16 == 0) +case 11: { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -lean_inc(x_3); -x_17 = lean_mk_array(x_5, x_3); -x_18 = lean_array_push(x_4, x_17); -x_19 = lean_box(0); -x_20 = l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform___lambda__1(x_3, x_1, x_2, x_19, x_18); -return x_20; +lean_object* x_33; +x_33 = lean_ctor_get(x_2, 2); +lean_inc(x_33); +lean_dec(x_2); +x_2 = x_33; +goto _start; } -else +default: { -lean_object* x_21; lean_object* x_22; -lean_dec(x_5); -x_21 = lean_box(0); -x_22 = l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform___lambda__1(x_3, x_1, x_2, x_21, x_4); -return x_22; +uint8_t x_35; +lean_dec(x_2); +x_35 = 0; +return x_35; } } } } } -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__2___lambda__1___closed__1() { _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_dec(x_3); -x_6 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_7 = l_Array_anyMUnsafe_any___at_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform___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_object* x_1; +x_1 = lean_mk_string_from_bytes("WellFoundedRelation", 19); +return x_1; } } -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__2___lambda__1___closed__2() { _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_dec(x_3); -x_6 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_7 = l_Array_anyMUnsafe_any___at_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform___spec__2(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_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__2___lambda__1___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_WF_GuessLex_generateCombinations_x3f_go___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5) { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__2___lambda__1___closed__3() { _start: { -uint8_t x_6; -x_6 = lean_usize_dec_eq(x_4, x_5); -if (x_6 == 0) +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("sizeOfWFRel", 11); +return x_1; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__2___lambda__1___closed__4() { +_start: { -lean_object* x_7; lean_object* x_8; uint8_t x_9; -x_7 = lean_array_uget(x_3, x_4); -x_8 = lean_unsigned_to_nat(0u); -x_9 = lean_nat_dec_lt(x_8, x_2); -if (x_9 == 0) +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_Elab_WF_GuessLex_getSizeOfParams___spec__2___lambda__1___closed__3; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getSizeOfParams___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) { +_start: { -lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l_instInhabitedNat; -x_11 = l___private_Init_Util_0__outOfBounds___rarg(x_10); -x_12 = lean_nat_dec_eq(x_7, x_11); -lean_dec(x_11); -lean_dec(x_7); -if (x_12 == 0) +lean_object* x_10; +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_1); +x_10 = l_Lean_Meta_getLevel(x_1, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_10) == 0) { -uint8_t x_13; -x_13 = 1; -return x_13; +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_11 = lean_ctor_get(x_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_box(0); +x_14 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_14, 0, x_11); +lean_ctor_set(x_14, 1, x_13); +x_15 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__2___lambda__1___closed__2; +lean_inc(x_14); +x_16 = l_Lean_Expr_const___override(x_15, x_14); +lean_inc(x_1); +x_17 = l_Lean_Expr_app___override(x_16, x_1); +x_18 = lean_box(0); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_19 = l_Lean_Meta_synthInstance(x_17, x_18, x_5, x_6, x_7, x_8, x_12); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +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_Elab_WF_GuessLex_mkSizeOf___closed__2; +lean_inc(x_14); +x_23 = l_Lean_Expr_const___override(x_22, x_14); +lean_inc(x_1); +x_24 = l_Lean_Expr_app___override(x_23, x_1); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_25 = l_Lean_Meta_synthInstance(x_24, x_18, x_5, x_6, x_7, x_8, 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; lean_object* x_30; lean_object* x_31; +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_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__2___lambda__1___closed__4; +x_29 = l_Lean_Expr_const___override(x_28, x_14); +x_30 = l_Lean_mkAppB(x_29, x_1, x_26); +x_31 = l_Lean_Meta_isExprDefEq(x_20, x_30, x_5, x_6, x_7, x_8, x_27); +if (lean_obj_tag(x_31) == 0) +{ +lean_object* x_32; uint8_t x_33; +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +x_33 = lean_unbox(x_32); +lean_dec(x_32); +if (x_33 == 0) +{ +uint8_t x_34; +lean_dec(x_2); +x_34 = !lean_is_exclusive(x_31); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_31, 0); +lean_dec(x_35); +x_36 = lean_box(0); +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_3); +lean_ctor_set(x_31, 0, x_37); +return x_31; } else { -size_t x_14; size_t x_15; -x_14 = 1; -x_15 = lean_usize_add(x_4, x_14); -x_4 = x_15; -goto _start; +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_38 = lean_ctor_get(x_31, 1); +lean_inc(x_38); +lean_dec(x_31); +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_3); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_38); +return x_41; } } else { -lean_object* x_17; uint8_t x_18; -x_17 = lean_array_fget(x_1, x_8); -x_18 = lean_nat_dec_eq(x_7, x_17); -lean_dec(x_17); -lean_dec(x_7); -if (x_18 == 0) +uint8_t x_42; +x_42 = !lean_is_exclusive(x_31); +if (x_42 == 0) { -uint8_t x_19; -x_19 = 1; -return x_19; +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_43 = lean_ctor_get(x_31, 0); +lean_dec(x_43); +x_44 = lean_array_push(x_3, 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_44); +lean_ctor_set(x_31, 0, x_46); +return x_31; } else { -size_t x_20; size_t x_21; -x_20 = 1; -x_21 = lean_usize_add(x_4, x_20); -x_4 = x_21; -goto _start; +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_47 = lean_ctor_get(x_31, 1); +lean_inc(x_47); +lean_dec(x_31); +x_48 = lean_array_push(x_3, x_2); +x_49 = lean_box(0); +x_50 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_50, 0, x_49); +lean_ctor_set(x_50, 1, x_48); +x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_47); +return x_51; } } } else { -uint8_t x_23; -x_23 = 0; -return x_23; +uint8_t x_52; +lean_dec(x_3); +lean_dec(x_2); +x_52 = !lean_is_exclusive(x_31); +if (x_52 == 0) +{ +return x_31; } +else +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_31, 0); +x_54 = lean_ctor_get(x_31, 1); +lean_inc(x_54); +lean_inc(x_53); +lean_dec(x_31); +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_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_generateCombinations_x3f_go___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -uint8_t x_12; -x_12 = lean_nat_dec_le(x_7, x_6); -if (x_12 == 0) -{ -lean_object* x_13; uint8_t x_14; -x_13 = lean_unsigned_to_nat(0u); -x_14 = lean_nat_dec_eq(x_5, x_13); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; uint8_t x_17; -lean_dec(x_9); -x_15 = lean_unsigned_to_nat(1u); -x_16 = lean_nat_sub(x_5, x_15); -lean_dec(x_5); -x_17 = l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_isForbidden(x_1, x_4, x_6); -if (x_17 == 0) -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_18 = lean_nat_add(x_4, x_15); -lean_inc(x_6); -lean_inc(x_10); -x_19 = lean_array_push(x_10, x_6); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_20 = l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_go(x_1, x_2, x_3, x_18, x_19, x_11); -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -if (lean_obj_tag(x_21) == 0) +} +else { -uint8_t x_22; -lean_dec(x_16); -lean_dec(x_10); +uint8_t x_56; +lean_dec(x_20); +lean_dec(x_14); +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_22 = !lean_is_exclusive(x_20); -if (x_22 == 0) +x_56 = !lean_is_exclusive(x_25); +if (x_56 == 0) { -lean_object* x_23; lean_object* x_24; -x_23 = lean_ctor_get(x_20, 0); -lean_dec(x_23); -x_24 = lean_box(0); -lean_ctor_set(x_20, 0, x_24); -return x_20; +return x_25; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_20, 1); -lean_inc(x_25); -lean_dec(x_20); -x_26 = lean_box(0); -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_26); -lean_ctor_set(x_27, 1, x_25); -return x_27; +lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_57 = lean_ctor_get(x_25, 0); +x_58 = lean_ctor_get(x_25, 1); +lean_inc(x_58); +lean_inc(x_57); +lean_dec(x_25); +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_28; lean_object* x_29; lean_object* x_30; -lean_dec(x_21); -x_28 = lean_ctor_get(x_20, 1); -lean_inc(x_28); -lean_dec(x_20); -x_29 = lean_nat_add(x_6, x_8); +uint8_t x_60; +lean_dec(x_14); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); -x_30 = lean_box(0); -x_5 = x_16; -x_6 = x_29; -x_9 = x_30; -x_11 = x_28; -goto _start; -} +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_60 = !lean_is_exclusive(x_19); +if (x_60 == 0) +{ +return x_19; } else { -lean_object* x_32; lean_object* x_33; -x_32 = lean_nat_add(x_6, x_8); -lean_dec(x_6); -x_33 = lean_box(0); -x_5 = x_16; -x_6 = x_32; -x_9 = x_33; -goto _start; +lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_61 = lean_ctor_get(x_19, 0); +x_62 = lean_ctor_get(x_19, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_19); +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; +} } } else { -lean_object* x_35; lean_object* x_36; -lean_dec(x_10); +uint8_t x_64; +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_35 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_35, 0, x_9); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_11); -return x_36; -} +x_64 = !lean_is_exclusive(x_10); +if (x_64 == 0) +{ +return x_10; } else { -lean_object* x_37; lean_object* x_38; +lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_65 = lean_ctor_get(x_10, 0); +x_66 = lean_ctor_get(x_10, 1); +lean_inc(x_66); +lean_inc(x_65); lean_dec(x_10); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_37 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_37, 0, x_9); -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; +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_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_go___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_Elab_WF_GuessLex_getSizeOfParams___spec__2(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_5; uint8_t x_6; -x_5 = lean_array_get_size(x_4); -x_6 = lean_nat_dec_lt(x_1, x_5); -lean_dec(x_5); -if (x_6 == 0) +uint8_t x_11; +x_11 = lean_usize_dec_lt(x_4, x_3); +if (x_11 == 0) { -lean_object* x_7; lean_object* x_8; -x_7 = l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform___closed__1; -x_8 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_8, 0, x_7); -lean_ctor_set(x_8, 1, x_4); -return x_8; +lean_object* x_12; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_5); +lean_ctor_set(x_12, 1, x_10); +return x_12; } else { -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_4); -return x_10; +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_13 = lean_array_uget(x_2, x_4); +x_23 = lean_ctor_get(x_5, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_5, 1); +lean_inc(x_24); +if (lean_is_exclusive(x_5)) { + lean_ctor_release(x_5, 0); + lean_ctor_release(x_5, 1); + x_25 = x_5; +} else { + lean_dec_ref(x_5); + x_25 = lean_box(0); } +x_26 = lean_ctor_get(x_23, 0); +lean_inc(x_26); +x_27 = lean_ctor_get(x_23, 1); +lean_inc(x_27); +x_28 = lean_ctor_get(x_23, 2); +lean_inc(x_28); +x_29 = lean_nat_dec_lt(x_26, x_27); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_13); +if (lean_is_scalar(x_25)) { + x_30 = lean_alloc_ctor(0, 2, 0); +} else { + x_30 = x_25; } +lean_ctor_set(x_30, 0, x_23); +lean_ctor_set(x_30, 1, x_24); +x_31 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_31, 0, x_30); +x_14 = x_31; +x_15 = x_10; +goto block_22; } -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_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) { -_start: +else { -lean_object* x_7; uint8_t x_8; -x_7 = lean_array_get_size(x_2); -x_8 = lean_nat_dec_lt(x_4, x_7); -lean_dec(x_7); -if (x_8 == 0) +uint8_t x_32; +x_32 = !lean_is_exclusive(x_23); +if (x_32 == 0) { -lean_object* x_9; lean_object* x_10; uint8_t x_11; -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_9 = lean_array_get_size(x_5); -x_10 = lean_unsigned_to_nat(0u); -x_11 = lean_nat_dec_lt(x_10, x_9); -if (x_11 == 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_49; +x_33 = lean_ctor_get(x_23, 2); +lean_dec(x_33); +x_34 = lean_ctor_get(x_23, 1); +lean_dec(x_34); +x_35 = lean_ctor_get(x_23, 0); +lean_dec(x_35); +x_36 = lean_nat_add(x_26, x_28); +lean_ctor_set(x_23, 0, x_36); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_49 = lean_infer_type(x_13, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_49) == 0) { -lean_object* x_12; lean_object* x_13; -lean_dec(x_9); -x_12 = lean_box(0); -x_13 = l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_go___lambda__1(x_3, x_12, x_5, x_6); -lean_dec(x_5); -lean_dec(x_3); -return x_13; +lean_object* x_50; lean_object* x_51; uint8_t x_52; +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); +lean_inc(x_50); +x_52 = l_Lean_Expr_hasAnyFVar_visit___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__1(x_1, x_50); +if (x_52 == 0) +{ +lean_object* x_53; lean_object* x_54; +x_53 = lean_box(0); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_24); +x_54 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__2___lambda__1(x_50, x_26, x_24, x_53, x_6, x_7, x_8, x_9, x_51); +if (lean_obj_tag(x_54) == 0) +{ +lean_object* x_55; lean_object* x_56; +lean_dec(x_24); +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_37 = x_55; +x_38 = x_56; +goto block_48; } else { -size_t x_14; size_t x_15; uint8_t x_16; -x_14 = 0; -x_15 = lean_usize_of_nat(x_9); -x_16 = l_Array_anyMUnsafe_any___at_Lean_Elab_WF_GuessLex_generateCombinations_x3f_go___spec__1(x_5, x_9, x_5, x_14, x_15); -lean_dec(x_9); -if (x_16 == 0) +uint8_t x_57; +x_57 = !lean_is_exclusive(x_54); +if (x_57 == 0) { -lean_object* x_17; lean_object* x_18; -x_17 = lean_box(0); -x_18 = l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_go___lambda__1(x_3, x_17, x_5, x_6); -lean_dec(x_5); -lean_dec(x_3); -return x_18; +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_Lean_Exception_isRuntime(x_58); +if (x_60 == 0) +{ +lean_object* x_61; +lean_free_object(x_54); +lean_dec(x_58); +x_61 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_61, 0, x_53); +lean_ctor_set(x_61, 1, x_24); +x_37 = x_61; +x_38 = x_59; +goto block_48; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -lean_inc(x_5); -x_19 = lean_array_push(x_6, x_5); -x_20 = lean_box(0); -x_21 = l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_go___lambda__1(x_3, x_20, x_5, x_19); -lean_dec(x_5); -lean_dec(x_3); -return x_21; +uint8_t x_62; +x_62 = lean_ctor_get_uint8(x_8, sizeof(void*)*11); +if (x_62 == 0) +{ +lean_dec(x_23); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +return x_54; +} +else +{ +lean_object* x_63; +lean_free_object(x_54); +lean_dec(x_58); +x_63 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_63, 0, x_53); +lean_ctor_set(x_63, 1, x_24); +x_37 = x_63; +x_38 = x_59; +goto block_48; } } } 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_array_fget(x_2, x_4); -x_23 = lean_unsigned_to_nat(0u); -x_24 = lean_unsigned_to_nat(1u); -x_25 = lean_box(0); -lean_inc(x_22); -x_26 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_generateCombinations_x3f_go___spec__2(x_1, x_2, x_3, x_4, x_22, x_23, x_22, x_24, x_25, x_5, x_6); -lean_dec(x_22); -lean_dec(x_4); -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -if (lean_obj_tag(x_27) == 0) +lean_object* x_64; lean_object* x_65; uint8_t x_66; +x_64 = lean_ctor_get(x_54, 0); +x_65 = lean_ctor_get(x_54, 1); +lean_inc(x_65); +lean_inc(x_64); +lean_dec(x_54); +x_66 = l_Lean_Exception_isRuntime(x_64); +if (x_66 == 0) { -uint8_t x_28; -x_28 = !lean_is_exclusive(x_26); -if (x_28 == 0) +lean_object* x_67; +lean_dec(x_64); +x_67 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_67, 0, x_53); +lean_ctor_set(x_67, 1, x_24); +x_37 = x_67; +x_38 = x_65; +goto block_48; +} +else { -lean_object* x_29; lean_object* x_30; -x_29 = lean_ctor_get(x_26, 0); -lean_dec(x_29); -x_30 = lean_box(0); -lean_ctor_set(x_26, 0, x_30); -return x_26; +uint8_t x_68; +x_68 = lean_ctor_get_uint8(x_8, sizeof(void*)*11); +if (x_68 == 0) +{ +lean_object* x_69; +lean_dec(x_23); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +x_69 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_69, 0, x_64); +lean_ctor_set(x_69, 1, x_65); +return x_69; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_26, 1); -lean_inc(x_31); +lean_object* x_70; +lean_dec(x_64); +x_70 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_70, 0, x_53); +lean_ctor_set(x_70, 1, x_24); +x_37 = x_70; +x_38 = x_65; +goto block_48; +} +} +} +} +} +else +{ +lean_object* x_71; +lean_dec(x_50); lean_dec(x_26); -x_32 = lean_box(0); -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; +x_71 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_71, 0, x_24); +x_37 = x_71; +x_38 = x_51; +goto block_48; } } else { -uint8_t x_34; -lean_dec(x_27); -x_34 = !lean_is_exclusive(x_26); -if (x_34 == 0) +uint8_t x_72; +lean_dec(x_26); +x_72 = !lean_is_exclusive(x_49); +if (x_72 == 0) { -lean_object* x_35; lean_object* x_36; -x_35 = lean_ctor_get(x_26, 0); -lean_dec(x_35); -x_36 = l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform___closed__1; -lean_ctor_set(x_26, 0, x_36); -return x_26; +lean_object* x_73; lean_object* x_74; uint8_t x_75; +x_73 = lean_ctor_get(x_49, 0); +x_74 = lean_ctor_get(x_49, 1); +x_75 = l_Lean_Exception_isRuntime(x_73); +if (x_75 == 0) +{ +lean_object* x_76; lean_object* x_77; +lean_free_object(x_49); +lean_dec(x_73); +x_76 = lean_box(0); +x_77 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_77, 0, x_76); +lean_ctor_set(x_77, 1, x_24); +x_37 = x_77; +x_38 = x_74; +goto block_48; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_26, 1); -lean_inc(x_37); -lean_dec(x_26); -x_38 = l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform___closed__1; -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_37); -return x_39; -} +uint8_t x_78; +x_78 = lean_ctor_get_uint8(x_8, sizeof(void*)*11); +if (x_78 == 0) +{ +lean_dec(x_23); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +return x_49; } +else +{ +lean_object* x_79; lean_object* x_80; +lean_free_object(x_49); +lean_dec(x_73); +x_79 = lean_box(0); +x_80 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_80, 0, x_79); +lean_ctor_set(x_80, 1, x_24); +x_37 = x_80; +x_38 = x_74; +goto block_48; } } } -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_WF_GuessLex_generateCombinations_x3f_go___spec__1___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; uint8_t x_8; lean_object* x_9; -x_6 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_7 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_8 = l_Array_anyMUnsafe_any___at_Lean_Elab_WF_GuessLex_generateCombinations_x3f_go___spec__1(x_1, x_2, x_3, x_6, x_7); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_9 = lean_box(x_8); -return x_9; -} +lean_object* x_81; lean_object* x_82; uint8_t x_83; +x_81 = lean_ctor_get(x_49, 0); +x_82 = lean_ctor_get(x_49, 1); +lean_inc(x_82); +lean_inc(x_81); +lean_dec(x_49); +x_83 = l_Lean_Exception_isRuntime(x_81); +if (x_83 == 0) +{ +lean_object* x_84; lean_object* x_85; +lean_dec(x_81); +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_24); +x_37 = x_85; +x_38 = x_82; +goto block_48; } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_generateCombinations_x3f_go___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) { -_start: +else { -lean_object* x_12; -x_12 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_generateCombinations_x3f_go___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +uint8_t x_86; +x_86 = lean_ctor_get_uint8(x_8, sizeof(void*)*11); +if (x_86 == 0) +{ +lean_object* x_87; +lean_dec(x_23); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_4); -return x_12; -} +lean_dec(x_6); +x_87 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_87, 0, x_81); +lean_ctor_set(x_87, 1, x_82); +return x_87; } -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_go___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: +else { -lean_object* x_5; -x_5 = l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_go___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_object* x_88; lean_object* x_89; +lean_dec(x_81); +x_88 = lean_box(0); +x_89 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_89, 0, x_88); +lean_ctor_set(x_89, 1, x_24); +x_37 = x_89; +x_38 = x_82; +goto block_48; } } -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_generateCombinations_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +} +} +block_48: { -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_4 = lean_unsigned_to_nat(0u); -x_5 = l_Lean_Elab_WF_GuessLex_naryVarNames___closed__1; -lean_inc(x_2); -lean_inc(x_1); -x_6 = l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform(x_1, x_2, x_4, x_5); -x_7 = lean_ctor_get(x_6, 0); -lean_inc(x_7); -x_8 = !lean_is_exclusive(x_7); -if (x_8 == 0) +switch (lean_obj_tag(x_37)) { +case 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); -lean_dec(x_9); -x_10 = lean_ctor_get(x_6, 1); -lean_inc(x_10); -lean_dec(x_6); -x_11 = l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_go(x_1, x_2, x_3, x_4, x_5, x_10); -x_12 = lean_ctor_get(x_11, 1); -lean_inc(x_12); -lean_dec(x_11); -lean_ctor_set(x_7, 0, x_12); -return x_7; +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_37, 1); +lean_inc(x_39); +lean_dec(x_37); +if (lean_is_scalar(x_25)) { + x_40 = lean_alloc_ctor(0, 2, 0); +} else { + x_40 = x_25; } -else +lean_ctor_set(x_40, 0, x_23); +lean_ctor_set(x_40, 1, x_39); +x_41 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_41, 0, x_40); +x_14 = x_41; +x_15 = x_38; +goto block_22; +} +case 1: { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -lean_dec(x_7); -x_13 = lean_ctor_get(x_6, 1); -lean_inc(x_13); -lean_dec(x_6); -x_14 = l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_go(x_1, x_2, x_3, x_4, x_5, x_13); -x_15 = lean_ctor_get(x_14, 1); -lean_inc(x_15); -lean_dec(x_14); -x_16 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_16, 0, x_15); -return x_16; +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_37, 0); +lean_inc(x_42); +lean_dec(x_37); +if (lean_is_scalar(x_25)) { + x_43 = lean_alloc_ctor(0, 2, 0); +} else { + x_43 = x_25; } +lean_ctor_set(x_43, 0, x_23); +lean_ctor_set(x_43, 1, x_42); +x_44 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_44, 0, x_43); +x_14 = x_44; +x_15 = x_38; +goto block_22; +} +default: +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_37, 0); +lean_inc(x_45); +lean_dec(x_37); +if (lean_is_scalar(x_25)) { + x_46 = lean_alloc_ctor(0, 2, 0); +} else { + x_46 = x_25; } +lean_ctor_set(x_46, 0, x_23); +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_14 = x_47; +x_15 = x_38; +goto block_22; } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_WF_GuessLex_generateMeasures___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: +} +} +} +else { -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_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_104; +lean_dec(x_23); +x_90 = lean_nat_add(x_26, x_28); +x_91 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_91, 0, x_90); +lean_ctor_set(x_91, 1, x_27); +lean_ctor_set(x_91, 2, x_28); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_104 = lean_infer_type(x_13, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_104) == 0) { -lean_object* x_10; lean_object* x_11; -x_10 = lean_ctor_get(x_8, 0); +lean_object* x_105; lean_object* x_106; uint8_t x_107; +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); +lean_inc(x_105); +x_107 = l_Lean_Expr_hasAnyFVar_visit___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__1(x_1, x_105); +if (x_107 == 0) +{ +lean_object* x_108; lean_object* x_109; +x_108 = lean_box(0); +lean_inc(x_9); +lean_inc(x_8); 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; +lean_inc(x_6); +lean_inc(x_24); +x_109 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__2___lambda__1(x_105, x_26, x_24, x_108, x_6, x_7, x_8, x_9, x_106); +if (lean_obj_tag(x_109) == 0) +{ +lean_object* x_110; lean_object* x_111; +lean_dec(x_24); +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_92 = x_110; +x_93 = x_111; +goto block_103; } 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_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; -} +lean_object* x_112; lean_object* x_113; lean_object* x_114; uint8_t x_115; +x_112 = lean_ctor_get(x_109, 0); +lean_inc(x_112); +x_113 = lean_ctor_get(x_109, 1); +lean_inc(x_113); +if (lean_is_exclusive(x_109)) { + lean_ctor_release(x_109, 0); + lean_ctor_release(x_109, 1); + x_114 = x_109; +} else { + lean_dec_ref(x_109); + x_114 = lean_box(0); } +x_115 = l_Lean_Exception_isRuntime(x_112); +if (x_115 == 0) +{ +lean_object* x_116; +lean_dec(x_114); +lean_dec(x_112); +x_116 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_116, 0, x_108); +lean_ctor_set(x_116, 1, x_24); +x_92 = x_116; +x_93 = x_113; +goto block_103; } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_generateMeasures___spec__2(size_t x_1, size_t x_2, lean_object* x_3) { -_start: +else { -uint8_t x_4; -x_4 = lean_usize_dec_lt(x_2, x_1); -if (x_4 == 0) +uint8_t x_117; +x_117 = lean_ctor_get_uint8(x_8, sizeof(void*)*11); +if (x_117 == 0) { -return x_3; +lean_object* x_118; +lean_dec(x_91); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +if (lean_is_scalar(x_114)) { + x_118 = lean_alloc_ctor(1, 2, 0); +} else { + x_118 = x_114; +} +lean_ctor_set(x_118, 0, x_112); +lean_ctor_set(x_118, 1, x_113); +return x_118; } else { -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; -x_5 = lean_array_uget(x_3, x_2); -x_6 = lean_unsigned_to_nat(0u); -x_7 = lean_array_uset(x_3, x_2, x_6); -x_8 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_8, 0, x_5); -x_9 = 1; -x_10 = lean_usize_add(x_2, x_9); -x_11 = lean_array_uset(x_7, x_2, x_8); -x_2 = x_10; -x_3 = x_11; -goto _start; +lean_object* x_119; +lean_dec(x_114); +lean_dec(x_112); +x_119 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_119, 0, x_108); +lean_ctor_set(x_119, 1, x_24); +x_92 = x_119; +x_93 = x_113; +goto block_103; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_generateMeasures___spec__3(size_t x_1, size_t x_2, lean_object* x_3) { -_start: -{ -uint8_t x_4; -x_4 = lean_usize_dec_lt(x_2, x_1); -if (x_4 == 0) -{ -return x_3; } else { -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; -x_5 = lean_array_uget(x_3, x_2); -x_6 = lean_unsigned_to_nat(0u); -x_7 = lean_array_uset(x_3, x_2, x_6); -x_8 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_8, 0, x_5); -x_9 = 1; -x_10 = lean_usize_add(x_2, x_9); -x_11 = lean_array_uset(x_7, x_2, x_8); -x_2 = x_10; -x_3 = x_11; -goto _start; +lean_object* x_120; +lean_dec(x_105); +lean_dec(x_26); +x_120 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_120, 0, x_24); +x_92 = x_120; +x_93 = x_106; +goto block_103; } } +else +{ +lean_object* x_121; lean_object* x_122; lean_object* x_123; uint8_t x_124; +lean_dec(x_26); +x_121 = lean_ctor_get(x_104, 0); +lean_inc(x_121); +x_122 = lean_ctor_get(x_104, 1); +lean_inc(x_122); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_123 = x_104; +} else { + lean_dec_ref(x_104); + x_123 = lean_box(0); +} +x_124 = l_Lean_Exception_isRuntime(x_121); +if (x_124 == 0) +{ +lean_object* x_125; lean_object* x_126; +lean_dec(x_123); +lean_dec(x_121); +x_125 = lean_box(0); +x_126 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_126, 0, x_125); +lean_ctor_set(x_126, 1, x_24); +x_92 = x_126; +x_93 = x_122; +goto block_103; } -static lean_object* _init_l_Lean_Elab_WF_GuessLex_generateMeasures___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Too many combinations", 21); -return x_1; +uint8_t x_127; +x_127 = lean_ctor_get_uint8(x_8, sizeof(void*)*11); +if (x_127 == 0) +{ +lean_object* x_128; +lean_dec(x_91); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +if (lean_is_scalar(x_123)) { + x_128 = lean_alloc_ctor(1, 2, 0); +} else { + x_128 = x_123; } +lean_ctor_set(x_128, 0, x_121); +lean_ctor_set(x_128, 1, x_122); +return x_128; } -static lean_object* _init_l_Lean_Elab_WF_GuessLex_generateMeasures___closed__2() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_WF_GuessLex_generateMeasures___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_129; lean_object* x_130; +lean_dec(x_123); +lean_dec(x_121); +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_24); +x_92 = x_130; +x_93 = x_122; +goto block_103; } } -static lean_object* _init_l_Lean_Elab_WF_GuessLex_generateMeasures___closed__3() { -_start: +} +block_103: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_WF_GuessLex_naryVarNames___closed__1; -x_2 = lean_array_get_size(x_1); -return x_2; +switch (lean_obj_tag(x_92)) { +case 0: +{ +lean_object* x_94; lean_object* x_95; lean_object* x_96; +x_94 = lean_ctor_get(x_92, 1); +lean_inc(x_94); +lean_dec(x_92); +if (lean_is_scalar(x_25)) { + x_95 = lean_alloc_ctor(0, 2, 0); +} else { + x_95 = x_25; +} +lean_ctor_set(x_95, 0, x_91); +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_14 = x_96; +x_15 = x_93; +goto block_22; +} +case 1: +{ +lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_97 = lean_ctor_get(x_92, 0); +lean_inc(x_97); +lean_dec(x_92); +if (lean_is_scalar(x_25)) { + x_98 = lean_alloc_ctor(0, 2, 0); +} else { + x_98 = x_25; } +lean_ctor_set(x_98, 0, x_91); +lean_ctor_set(x_98, 1, x_97); +x_99 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_99, 0, x_98); +x_14 = x_99; +x_15 = x_93; +goto block_22; } -static size_t _init_l_Lean_Elab_WF_GuessLex_generateMeasures___closed__4() { -_start: +default: { -lean_object* x_1; size_t x_2; -x_1 = l_Lean_Elab_WF_GuessLex_generateMeasures___closed__3; -x_2 = lean_usize_of_nat(x_1); -return x_2; +lean_object* x_100; lean_object* x_101; lean_object* x_102; +x_100 = lean_ctor_get(x_92, 0); +lean_inc(x_100); +lean_dec(x_92); +if (lean_is_scalar(x_25)) { + x_101 = lean_alloc_ctor(0, 2, 0); +} else { + x_101 = x_25; } +lean_ctor_set(x_101, 0, x_91); +lean_ctor_set(x_101, 1, x_100); +x_102 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_102, 0, x_101); +x_14 = x_102; +x_15 = x_93; +goto block_22; } -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_generateMeasures(lean_object* x_1, 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: +} +} +} +} +block_22: { -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; size_t x_15; size_t x_16; lean_object* x_17; -x_8 = lean_unsigned_to_nat(32u); -lean_inc(x_2); -x_9 = l_Lean_Elab_WF_GuessLex_generateCombinations_x3f(x_1, x_2, x_8); -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_16; lean_object* x_17; lean_dec(x_9); -x_11 = lean_array_get_size(x_2); -lean_dec(x_2); -x_12 = lean_unsigned_to_nat(1u); -x_13 = lean_nat_dec_lt(x_12, x_11); -x_14 = lean_array_get_size(x_10); -x_15 = lean_usize_of_nat(x_14); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +x_16 = lean_ctor_get(x_14, 0); +lean_inc(x_16); lean_dec(x_14); -x_16 = 0; -x_17 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_generateMeasures___spec__2(x_15, x_16, x_10); -if (x_13 == 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 { -size_t x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +lean_object* x_18; size_t x_19; size_t x_20; +x_18 = lean_ctor_get(x_14, 0); +lean_inc(x_18); +lean_dec(x_14); +x_19 = 1; +x_20 = lean_usize_add(x_4, x_19); +x_4 = x_20; +x_5 = x_18; +x_10 = x_15; +goto _start; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_getSizeOfParams___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_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; +x_9 = lean_array_get_size(x_2); +x_10 = l_Array_toSubarray___rarg(x_2, x_1, x_9); +x_11 = l_Array_ofSubarray___rarg(x_10); +x_12 = lean_array_get_size(x_11); +x_13 = lean_unsigned_to_nat(0u); +x_14 = lean_unsigned_to_nat(1u); +lean_inc(x_12); +x_15 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_15, 0, x_13); +lean_ctor_set(x_15, 1, x_12); +lean_ctor_set(x_15, 2, x_14); +x_16 = l_Lean_Elab_WF_GuessLex_naryVarNames___closed__1; +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set(x_17, 1, x_16); +x_18 = lean_usize_of_nat(x_12); +lean_dec(x_12); +x_19 = 0; +x_20 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__2(x_11, x_11, x_18, x_19, x_17, x_4, x_5, x_6, x_7, x_8); lean_dec(x_11); -x_18 = l_Lean_Elab_WF_GuessLex_generateMeasures___closed__4; -x_19 = l_Lean_Elab_WF_GuessLex_naryVarNames___closed__1; -x_20 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_generateMeasures___spec__3(x_18, x_16, x_19); -x_21 = l_Array_append___rarg(x_17, 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_7); -return x_22; +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 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +lean_dec(x_22); +lean_ctor_set(x_20, 0, x_23); +return x_20; } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; size_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_23 = l_List_range(x_11); -x_24 = l_List_redLength___rarg(x_23); -x_25 = lean_mk_empty_array_with_capacity(x_24); +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_24 = lean_ctor_get(x_20, 0); +x_25 = lean_ctor_get(x_20, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_20); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); lean_dec(x_24); -x_26 = l_List_toArrayAux___rarg(x_23, x_25); -x_27 = lean_array_get_size(x_26); -x_28 = lean_usize_of_nat(x_27); -lean_dec(x_27); -x_29 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_generateMeasures___spec__3(x_28, x_16, x_26); -x_30 = l_Array_append___rarg(x_17, 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_7); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_25); +return x_27; +} +} +else +{ +uint8_t x_28; +x_28 = !lean_is_exclusive(x_20); +if (x_28 == 0) +{ +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 = 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_throwError___at_Lean_Elab_WF_GuessLex_generateMeasures___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +} +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_getSizeOfParams(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_7; -x_7 = l_Lean_throwError___at_Lean_Elab_WF_GuessLex_generateMeasures___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_object* x_8; lean_object* x_9; uint8_t x_10; lean_object* x_11; +x_8 = lean_ctor_get(x_2, 5); +lean_inc(x_8); lean_dec(x_2); -return x_7; +x_9 = lean_alloc_closure((void*)(l_Lean_Elab_WF_GuessLex_getSizeOfParams___lambda__1___boxed), 8, 1); +lean_closure_set(x_9, 0, x_1); +x_10 = 0; +x_11 = l_Lean_Meta_lambdaTelescope___at_Lean_PrettyPrinter_Delaborator_returnsPi___spec__1___rarg(x_8, x_9, x_10, x_3, x_4, x_5, x_6, x_7); +return x_11; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_generateMeasures___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Expr_hasAnyFVar_visit___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { -size_t x_4; size_t x_5; lean_object* x_6; -x_4 = lean_unbox_usize(x_1); +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_Expr_hasAnyFVar_visit___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__1(x_1, x_2); lean_dec(x_1); -x_5 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_generateMeasures___spec__2(x_4, x_5, x_3); -return x_6; +x_4 = lean_box(x_3); +return x_4; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_generateMeasures___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__2___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -size_t x_4; size_t x_5; lean_object* x_6; -x_4 = lean_unbox_usize(x_1); -lean_dec(x_1); -x_5 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_generateMeasures___spec__3(x_4, x_5, x_3); -return x_6; +lean_object* x_10; +x_10 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__2___lambda__1(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_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_generateMeasures___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getSizeOfParams___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) { _start: { -lean_object* x_8; -x_8 = l_Lean_Elab_WF_GuessLex_generateMeasures(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); +size_t x_11; size_t x_12; lean_object* x_13; +x_11 = lean_unbox_usize(x_3); lean_dec(x_3); -return x_8; +x_12 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_13 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__2(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_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_getSizeOfParams___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_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_6 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_6, 0, x_1); -lean_ctor_set(x_6, 1, x_2); -x_7 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_7, 0, x_3); -lean_ctor_set(x_7, 1, x_6); -x_8 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_8, 0, x_7); -x_9 = lean_apply_2(x_4, lean_box(0), x_8); +lean_object* x_9; +x_9 = l_Lean_Elab_WF_GuessLex_getSizeOfParams___lambda__1(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_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getNatParams___spec__1___closed__1() { _start: { -uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_5 = 1; -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_1); -x_8 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_8, 0, x_2); -lean_ctor_set(x_8, 1, x_7); -x_9 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_9, 0, x_8); -x_10 = lean_apply_2(x_3, lean_box(0), x_9); -return x_10; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_WF_GuessLex_GuessLexRel_toNatRel___closed__7; +x_3 = l_Lean_Expr_const___override(x_2, x_1); +return x_3; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___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_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getNatParams___spec__1(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -uint8_t x_8; uint8_t x_9; -x_8 = 0; -x_9 = l_Lean_Elab_WF_GuessLex_instDecidableEqGuessLexRel(x_7, x_8); -if (x_9 == 0) +uint8_t x_11; +x_11 = lean_usize_dec_lt(x_4, x_3); +if (x_11 == 0) { -lean_object* x_10; uint8_t x_11; uint8_t x_12; -x_10 = lean_array_push(x_1, x_2); -x_11 = 2; -x_12 = l_Lean_Elab_WF_GuessLex_instDecidableEqGuessLexRel(x_7, x_11); -if (x_12 == 0) +lean_object* x_12; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_5); +lean_ctor_set(x_12, 1, x_10); +return x_12; +} +else { -uint8_t x_13; uint8_t x_14; -x_13 = 1; -x_14 = l_Lean_Elab_WF_GuessLex_instDecidableEqGuessLexRel(x_7, x_13); +lean_object* x_13; uint8_t x_14; +x_13 = lean_array_uget(x_2, x_4); +x_14 = !lean_is_exclusive(x_5); if (x_14 == 0) { -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; +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_15 = lean_ctor_get(x_5, 0); +x_16 = lean_ctor_get(x_5, 1); +x_17 = lean_ctor_get(x_15, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_15, 1); +lean_inc(x_18); +x_19 = lean_ctor_get(x_15, 2); +lean_inc(x_19); +x_20 = lean_nat_dec_lt(x_17, x_18); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_13); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -x_15 = lean_ctor_get(x_3, 0); -lean_inc(x_15); -lean_dec(x_3); -x_16 = lean_ctor_get(x_15, 1); -lean_inc(x_16); -lean_dec(x_15); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_4); -lean_ctor_set(x_17, 1, x_10); -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_17); -x_21 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_21, 0, x_20); -x_22 = lean_apply_2(x_16, lean_box(0), x_21); -return x_22; +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_5); +lean_ctor_set(x_21, 1, x_10); +return x_21; } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_23 = lean_ctor_get(x_3, 0); -lean_inc(x_23); -lean_dec(x_3); -x_24 = lean_ctor_get(x_23, 1); -lean_inc(x_24); +uint8_t x_22; +x_22 = !lean_is_exclusive(x_15); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_23 = lean_ctor_get(x_15, 2); lean_dec(x_23); -x_25 = lean_box(0); -lean_inc(x_24); -x_26 = lean_apply_2(x_24, lean_box(0), x_25); -x_27 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg___lambda__1___boxed), 5, 4); -lean_closure_set(x_27, 0, x_4); -lean_closure_set(x_27, 1, x_10); -lean_closure_set(x_27, 2, x_5); -lean_closure_set(x_27, 3, x_24); -x_28 = lean_apply_4(x_6, lean_box(0), lean_box(0), x_26, x_27); -return x_28; -} -} -else +x_24 = lean_ctor_get(x_15, 1); +lean_dec(x_24); +x_25 = lean_ctor_get(x_15, 0); +lean_dec(x_25); +x_26 = lean_nat_add(x_17, x_19); +lean_ctor_set(x_15, 0, x_26); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_27 = lean_infer_type(x_13, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_27) == 0) { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_29 = lean_ctor_get(x_3, 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; uint8_t x_36; +x_28 = lean_ctor_get(x_6, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_27, 0); lean_inc(x_29); -lean_dec(x_3); -x_30 = lean_ctor_get(x_29, 1); -lean_inc(x_30); -lean_dec(x_29); -x_31 = lean_box(0); +x_30 = lean_ctor_get(x_27, 1); lean_inc(x_30); -x_32 = lean_apply_2(x_30, lean_box(0), x_31); -x_33 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg___lambda__1___boxed), 5, 4); -lean_closure_set(x_33, 0, x_4); -lean_closure_set(x_33, 1, x_10); -lean_closure_set(x_33, 2, x_5); -lean_closure_set(x_33, 3, x_30); -x_34 = lean_apply_4(x_6, lean_box(0), lean_box(0), x_32, x_33); -return x_34; -} +lean_dec(x_27); +x_31 = lean_ctor_get(x_6, 1); +lean_inc(x_31); +x_32 = lean_ctor_get(x_6, 2); +lean_inc(x_32); +x_33 = lean_ctor_get(x_6, 3); +lean_inc(x_33); +x_34 = lean_ctor_get(x_6, 4); +lean_inc(x_34); +x_35 = lean_ctor_get(x_6, 5); +lean_inc(x_35); +x_36 = !lean_is_exclusive(x_28); +if (x_36 == 0) +{ +uint8_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_37 = 2; +lean_ctor_set_uint8(x_28, 9, x_37); +x_38 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_38, 0, x_28); +lean_ctor_set(x_38, 1, x_31); +lean_ctor_set(x_38, 2, x_32); +lean_ctor_set(x_38, 3, x_33); +lean_ctor_set(x_38, 4, x_34); +lean_ctor_set(x_38, 5, x_35); +x_39 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getNatParams___spec__1___closed__1; +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_40 = l_Lean_Meta_isExprDefEq(x_29, x_39, x_38, x_7, x_8, x_9, x_30); +if (lean_obj_tag(x_40) == 0) +{ +lean_object* x_41; uint8_t x_42; +x_41 = lean_ctor_get(x_40, 0); +lean_inc(x_41); +x_42 = lean_unbox(x_41); +lean_dec(x_41); +if (x_42 == 0) +{ +lean_object* x_43; size_t x_44; size_t x_45; +lean_dec(x_17); +x_43 = lean_ctor_get(x_40, 1); +lean_inc(x_43); +lean_dec(x_40); +x_44 = 1; +x_45 = lean_usize_add(x_4, x_44); +x_4 = x_45; +x_10 = x_43; +goto _start; } 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_dec(x_4); -lean_dec(x_2); -x_35 = lean_ctor_get(x_3, 0); -lean_inc(x_35); -lean_dec(x_3); -x_36 = lean_ctor_get(x_35, 1); -lean_inc(x_36); -lean_dec(x_35); -x_37 = lean_box(0); -lean_inc(x_36); -x_38 = lean_apply_2(x_36, lean_box(0), x_37); -x_39 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg___lambda__2___boxed), 4, 3); -lean_closure_set(x_39, 0, x_1); -lean_closure_set(x_39, 1, x_5); -lean_closure_set(x_39, 2, x_36); -x_40 = lean_apply_4(x_6, lean_box(0), lean_box(0), x_38, x_39); -return x_40; -} +lean_object* x_47; lean_object* x_48; size_t x_49; size_t x_50; +x_47 = lean_ctor_get(x_40, 1); +lean_inc(x_47); +lean_dec(x_40); +x_48 = lean_array_push(x_16, x_17); +lean_ctor_set(x_5, 1, x_48); +x_49 = 1; +x_50 = lean_usize_add(x_4, x_49); +x_4 = x_50; +x_10 = x_47; +goto _start; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg___lambda__4(lean_object* x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, size_t x_6, lean_object* x_7) { -_start: -{ -if (lean_obj_tag(x_7) == 0) +else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_8 = lean_ctor_get(x_7, 0); -lean_inc(x_8); -lean_dec(x_7); -x_9 = lean_ctor_get(x_1, 0); -lean_inc(x_9); -lean_dec(x_1); -x_10 = lean_ctor_get(x_9, 1); -lean_inc(x_10); +uint8_t x_52; +lean_dec(x_15); +lean_dec(x_17); +lean_free_object(x_5); +lean_dec(x_16); lean_dec(x_9); -x_11 = lean_apply_2(x_10, lean_box(0), x_8); -return x_11; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +x_52 = !lean_is_exclusive(x_40); +if (x_52 == 0) +{ +return x_40; } else { -lean_object* x_12; size_t x_13; size_t x_14; lean_object* x_15; -x_12 = lean_ctor_get(x_7, 0); -lean_inc(x_12); -lean_dec(x_7); -x_13 = 1; -x_14 = lean_usize_add(x_2, x_13); -x_15 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg(x_1, x_3, x_4, x_5, x_6, x_14, x_12); -return x_15; +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_40, 0); +x_54 = lean_ctor_get(x_40, 1); +lean_inc(x_54); +lean_inc(x_53); +lean_dec(x_40); +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_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7) { -_start: -{ -uint8_t x_8; -x_8 = lean_usize_dec_lt(x_6, x_5); -if (x_8 == 0) +else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_9 = lean_ctor_get(x_1, 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; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_56 = lean_ctor_get_uint8(x_28, 0); +x_57 = lean_ctor_get_uint8(x_28, 1); +x_58 = lean_ctor_get_uint8(x_28, 2); +x_59 = lean_ctor_get_uint8(x_28, 3); +x_60 = lean_ctor_get_uint8(x_28, 4); +x_61 = lean_ctor_get_uint8(x_28, 5); +x_62 = lean_ctor_get_uint8(x_28, 6); +x_63 = lean_ctor_get_uint8(x_28, 7); +x_64 = lean_ctor_get_uint8(x_28, 8); +x_65 = lean_ctor_get_uint8(x_28, 10); +x_66 = lean_ctor_get_uint8(x_28, 11); +lean_dec(x_28); +x_67 = 2; +x_68 = lean_alloc_ctor(0, 0, 12); +lean_ctor_set_uint8(x_68, 0, x_56); +lean_ctor_set_uint8(x_68, 1, x_57); +lean_ctor_set_uint8(x_68, 2, x_58); +lean_ctor_set_uint8(x_68, 3, x_59); +lean_ctor_set_uint8(x_68, 4, x_60); +lean_ctor_set_uint8(x_68, 5, x_61); +lean_ctor_set_uint8(x_68, 6, x_62); +lean_ctor_set_uint8(x_68, 7, x_63); +lean_ctor_set_uint8(x_68, 8, x_64); +lean_ctor_set_uint8(x_68, 9, x_67); +lean_ctor_set_uint8(x_68, 10, x_65); +lean_ctor_set_uint8(x_68, 11, x_66); +x_69 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_69, 0, x_68); +lean_ctor_set(x_69, 1, x_31); +lean_ctor_set(x_69, 2, x_32); +lean_ctor_set(x_69, 3, x_33); +lean_ctor_set(x_69, 4, x_34); +lean_ctor_set(x_69, 5, x_35); +x_70 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getNatParams___spec__1___closed__1; lean_inc(x_9); -lean_dec(x_1); -x_10 = lean_ctor_get(x_9, 1); -lean_inc(x_10); -lean_dec(x_9); -x_11 = lean_apply_2(x_10, lean_box(0), x_7); -return x_11; +lean_inc(x_8); +lean_inc(x_7); +x_71 = l_Lean_Meta_isExprDefEq(x_29, x_70, x_69, x_7, x_8, x_9, x_30); +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); +lean_dec(x_72); +if (x_73 == 0) +{ +lean_object* x_74; size_t x_75; size_t x_76; +lean_dec(x_17); +x_74 = lean_ctor_get(x_71, 1); +lean_inc(x_74); +lean_dec(x_71); +x_75 = 1; +x_76 = lean_usize_add(x_4, x_75); +x_4 = x_76; +x_10 = x_74; +goto _start; } else { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_12 = lean_array_uget(x_4, x_6); -x_13 = lean_ctor_get(x_7, 1); -lean_inc(x_13); -x_14 = lean_ctor_get(x_1, 1); -lean_inc(x_14); -x_15 = lean_ctor_get(x_7, 0); -lean_inc(x_15); -lean_dec(x_7); -x_16 = lean_ctor_get(x_13, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_13, 1); -lean_inc(x_17); -lean_dec(x_13); -x_18 = lean_array_uget(x_4, x_6); -lean_inc(x_3); -x_19 = lean_apply_1(x_18, x_3); -lean_inc(x_2); -lean_inc(x_1); -x_20 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg___lambda__3___boxed), 7, 6); -lean_closure_set(x_20, 0, x_17); -lean_closure_set(x_20, 1, x_12); -lean_closure_set(x_20, 2, x_1); -lean_closure_set(x_20, 3, x_16); -lean_closure_set(x_20, 4, x_15); -lean_closure_set(x_20, 5, x_2); -lean_inc(x_2); -x_21 = lean_apply_4(x_2, lean_box(0), lean_box(0), x_19, x_20); -x_22 = lean_box_usize(x_6); -x_23 = lean_box_usize(x_5); -x_24 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg___lambda__4___boxed), 7, 6); -lean_closure_set(x_24, 0, x_1); -lean_closure_set(x_24, 1, x_22); -lean_closure_set(x_24, 2, x_2); -lean_closure_set(x_24, 3, x_3); -lean_closure_set(x_24, 4, x_4); -lean_closure_set(x_24, 5, x_23); -x_25 = lean_apply_4(x_14, lean_box(0), lean_box(0), x_21, x_24); -return x_25; -} +lean_object* x_78; lean_object* x_79; size_t x_80; size_t x_81; +x_78 = lean_ctor_get(x_71, 1); +lean_inc(x_78); +lean_dec(x_71); +x_79 = lean_array_push(x_16, x_17); +lean_ctor_set(x_5, 1, x_79); +x_80 = 1; +x_81 = lean_usize_add(x_4, x_80); +x_4 = x_81; +x_10 = x_78; +goto _start; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1(lean_object* x_1, lean_object* x_2) { -_start: +else { -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg___boxed), 7, 0); -return x_3; -} +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +lean_dec(x_15); +lean_dec(x_17); +lean_free_object(x_5); +lean_dec(x_16); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +x_83 = lean_ctor_get(x_71, 0); +lean_inc(x_83); +x_84 = lean_ctor_get(x_71, 1); +lean_inc(x_84); +if (lean_is_exclusive(x_71)) { + lean_ctor_release(x_71, 0); + lean_ctor_release(x_71, 1); + x_85 = x_71; +} else { + lean_dec_ref(x_71); + x_85 = lean_box(0); } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2___rarg___lambda__1(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; -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 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_5, 0, x_2); -x_6 = lean_box(0); -x_7 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_7, 0, x_5); -lean_ctor_set(x_7, 1, x_6); -x_8 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_8, 0, x_7); -x_9 = lean_apply_2(x_4, lean_box(0), x_8); -return x_9; +if (lean_is_scalar(x_85)) { + x_86 = lean_alloc_ctor(1, 2, 0); +} else { + x_86 = x_85; } +lean_ctor_set(x_86, 0, x_83); +lean_ctor_set(x_86, 1, x_84); +return x_86; } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2___rarg___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_9 = l_Array_eraseIdx___rarg(x_1, x_2); -x_10 = lean_array_push(x_3, x_4); -lean_inc(x_5); -x_11 = l_Lean_Elab_WF_GuessLex_solve_go___rarg(x_5, x_9, x_6, x_10); -x_12 = lean_alloc_closure((void*)(l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2___rarg___lambda__1), 2, 1); -lean_closure_set(x_12, 0, x_5); -x_13 = lean_apply_4(x_7, lean_box(0), lean_box(0), x_11, x_12); -return x_13; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2___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) { -_start: -{ -lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_9 = lean_ctor_get(x_8, 1); -lean_inc(x_9); -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) +else { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +uint8_t x_87; +lean_dec(x_15); +lean_dec(x_17); +lean_free_object(x_5); +lean_dec(x_16); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -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(1, 1, 0); -lean_ctor_set(x_14, 0, x_2); -x_15 = lean_apply_2(x_13, lean_box(0), x_14); -return x_15; +x_87 = !lean_is_exclusive(x_27); +if (x_87 == 0) +{ +return x_27; } else { -lean_object* x_16; uint8_t x_17; -x_16 = lean_ctor_get(x_8, 0); -lean_inc(x_16); -lean_dec(x_8); -x_17 = lean_unbox(x_16); -lean_dec(x_16); -if (x_17 == 0) -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -lean_dec(x_9); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_18 = lean_ctor_get(x_1, 0); -lean_inc(x_18); -lean_dec(x_1); -x_19 = lean_ctor_get(x_18, 1); -lean_inc(x_19); -lean_dec(x_18); -x_20 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_20, 0, x_2); -x_21 = lean_apply_2(x_19, lean_box(0), x_20); -return x_21; +lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_88 = lean_ctor_get(x_27, 0); +x_89 = lean_ctor_get(x_27, 1); +lean_inc(x_89); +lean_inc(x_88); +lean_dec(x_27); +x_90 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_90, 0, x_88); +lean_ctor_set(x_90, 1, x_89); +return x_90; +} +} } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -lean_dec(x_2); -x_22 = lean_ctor_get(x_9, 1); -lean_inc(x_22); -lean_dec(x_9); +lean_object* x_91; lean_object* x_92; lean_object* x_93; +lean_dec(x_15); +x_91 = lean_nat_add(x_17, x_19); +x_92 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_92, 0, x_91); +lean_ctor_set(x_92, 1, x_18); +lean_ctor_set(x_92, 2, x_19); +lean_inc(x_9); +lean_inc(x_8); lean_inc(x_7); -lean_inc(x_1); -x_23 = lean_alloc_closure((void*)(l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2___rarg___lambda__2___boxed), 8, 7); -lean_closure_set(x_23, 0, x_3); -lean_closure_set(x_23, 1, x_4); -lean_closure_set(x_23, 2, x_5); -lean_closure_set(x_23, 3, x_6); -lean_closure_set(x_23, 4, x_1); -lean_closure_set(x_23, 5, x_22); -lean_closure_set(x_23, 6, x_7); -x_24 = lean_ctor_get(x_1, 0); -lean_inc(x_24); -lean_dec(x_1); -x_25 = lean_ctor_get(x_24, 1); -lean_inc(x_25); -lean_dec(x_24); -x_26 = lean_box(0); -x_27 = lean_apply_2(x_25, lean_box(0), x_26); -x_28 = lean_apply_4(x_7, lean_box(0), lean_box(0), x_27, x_23); -return x_28; +lean_inc(x_6); +x_93 = lean_infer_type(x_13, x_6, x_7, x_8, x_9, x_10); +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; lean_object* x_100; lean_object* x_101; uint8_t x_102; uint8_t x_103; 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; lean_object* x_113; uint8_t x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; +x_94 = lean_ctor_get(x_6, 0); +lean_inc(x_94); +x_95 = lean_ctor_get(x_93, 0); +lean_inc(x_95); +x_96 = lean_ctor_get(x_93, 1); +lean_inc(x_96); +lean_dec(x_93); +x_97 = lean_ctor_get(x_6, 1); +lean_inc(x_97); +x_98 = lean_ctor_get(x_6, 2); +lean_inc(x_98); +x_99 = lean_ctor_get(x_6, 3); +lean_inc(x_99); +x_100 = lean_ctor_get(x_6, 4); +lean_inc(x_100); +x_101 = lean_ctor_get(x_6, 5); +lean_inc(x_101); +x_102 = lean_ctor_get_uint8(x_94, 0); +x_103 = lean_ctor_get_uint8(x_94, 1); +x_104 = lean_ctor_get_uint8(x_94, 2); +x_105 = lean_ctor_get_uint8(x_94, 3); +x_106 = lean_ctor_get_uint8(x_94, 4); +x_107 = lean_ctor_get_uint8(x_94, 5); +x_108 = lean_ctor_get_uint8(x_94, 6); +x_109 = lean_ctor_get_uint8(x_94, 7); +x_110 = lean_ctor_get_uint8(x_94, 8); +x_111 = lean_ctor_get_uint8(x_94, 10); +x_112 = lean_ctor_get_uint8(x_94, 11); +if (lean_is_exclusive(x_94)) { + x_113 = x_94; +} else { + lean_dec_ref(x_94); + x_113 = lean_box(0); } +x_114 = 2; +if (lean_is_scalar(x_113)) { + x_115 = lean_alloc_ctor(0, 0, 12); +} else { + x_115 = x_113; +} +lean_ctor_set_uint8(x_115, 0, x_102); +lean_ctor_set_uint8(x_115, 1, x_103); +lean_ctor_set_uint8(x_115, 2, x_104); +lean_ctor_set_uint8(x_115, 3, x_105); +lean_ctor_set_uint8(x_115, 4, x_106); +lean_ctor_set_uint8(x_115, 5, x_107); +lean_ctor_set_uint8(x_115, 6, x_108); +lean_ctor_set_uint8(x_115, 7, x_109); +lean_ctor_set_uint8(x_115, 8, x_110); +lean_ctor_set_uint8(x_115, 9, x_114); +lean_ctor_set_uint8(x_115, 10, x_111); +lean_ctor_set_uint8(x_115, 11, x_112); +x_116 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_116, 0, x_115); +lean_ctor_set(x_116, 1, x_97); +lean_ctor_set(x_116, 2, x_98); +lean_ctor_set(x_116, 3, x_99); +lean_ctor_set(x_116, 4, x_100); +lean_ctor_set(x_116, 5, x_101); +x_117 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getNatParams___spec__1___closed__1; +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_118 = l_Lean_Meta_isExprDefEq(x_95, x_117, x_116, x_7, x_8, x_9, x_96); +if (lean_obj_tag(x_118) == 0) +{ +lean_object* x_119; uint8_t x_120; +x_119 = lean_ctor_get(x_118, 0); +lean_inc(x_119); +x_120 = lean_unbox(x_119); +lean_dec(x_119); +if (x_120 == 0) +{ +lean_object* x_121; size_t x_122; size_t x_123; +lean_dec(x_17); +x_121 = lean_ctor_get(x_118, 1); +lean_inc(x_121); +lean_dec(x_118); +lean_ctor_set(x_5, 0, x_92); +x_122 = 1; +x_123 = lean_usize_add(x_4, x_122); +x_4 = x_123; +x_10 = x_121; +goto _start; } +else +{ +lean_object* x_125; lean_object* x_126; size_t x_127; size_t x_128; +x_125 = lean_ctor_get(x_118, 1); +lean_inc(x_125); +lean_dec(x_118); +x_126 = lean_array_push(x_16, x_17); +lean_ctor_set(x_5, 1, x_126); +lean_ctor_set(x_5, 0, x_92); +x_127 = 1; +x_128 = lean_usize_add(x_4, x_127); +x_4 = x_128; +x_10 = x_125; +goto _start; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2___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) { -_start: -{ -if (lean_obj_tag(x_14) == 0) +else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); +lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; +lean_dec(x_92); +lean_dec(x_17); +lean_free_object(x_5); +lean_dec(x_16); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -lean_dec(x_14); -x_16 = lean_ctor_get(x_1, 0); -lean_inc(x_16); -lean_dec(x_1); -x_17 = lean_ctor_get(x_16, 1); -lean_inc(x_17); -lean_dec(x_16); -x_18 = lean_apply_2(x_17, lean_box(0), x_15); -return x_18; +x_130 = lean_ctor_get(x_118, 0); +lean_inc(x_130); +x_131 = lean_ctor_get(x_118, 1); +lean_inc(x_131); +if (lean_is_exclusive(x_118)) { + lean_ctor_release(x_118, 0); + lean_ctor_release(x_118, 1); + x_132 = x_118; +} else { + lean_dec_ref(x_118); + x_132 = lean_box(0); } -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_14, 0); -lean_inc(x_19); -lean_dec(x_14); -x_20 = lean_nat_add(x_2, x_3); -lean_dec(x_2); -x_21 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2___rarg(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_3, x_13, x_20, lean_box(0), x_19); -lean_dec(x_13); -return x_21; +if (lean_is_scalar(x_132)) { + x_133 = lean_alloc_ctor(1, 2, 0); +} else { + x_133 = x_132; } +lean_ctor_set(x_133, 0, x_130); +lean_ctor_set(x_133, 1, x_131); +return x_133; } } -static lean_object* _init_l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2___rarg___closed__1() { -_start: +else { -uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = 0; -x_2 = l_Lean_Elab_WF_GuessLex_naryVarNames___closed__1; -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_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; +lean_dec(x_92); +lean_dec(x_17); +lean_free_object(x_5); +lean_dec(x_16); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +x_134 = lean_ctor_get(x_93, 0); +lean_inc(x_134); +x_135 = lean_ctor_get(x_93, 1); +lean_inc(x_135); +if (lean_is_exclusive(x_93)) { + lean_ctor_release(x_93, 0); + lean_ctor_release(x_93, 1); + x_136 = x_93; +} else { + lean_dec_ref(x_93); + x_136 = lean_box(0); } +if (lean_is_scalar(x_136)) { + x_137 = lean_alloc_ctor(1, 2, 0); +} else { + x_137 = x_136; } -static lean_object* _init_l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2___rarg___closed__2() { -_start: -{ -uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = 1; -x_2 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2___rarg___closed__1; -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_ctor_set(x_137, 0, x_134); +lean_ctor_set(x_137, 1, x_135); +return x_137; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___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, lean_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; -x_16 = lean_nat_dec_lt(x_13, x_10); -if (x_16 == 0) +lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; uint8_t x_143; +x_138 = lean_ctor_get(x_5, 0); +x_139 = lean_ctor_get(x_5, 1); +lean_inc(x_139); +lean_inc(x_138); +lean_dec(x_5); +x_140 = lean_ctor_get(x_138, 0); +lean_inc(x_140); +x_141 = lean_ctor_get(x_138, 1); +lean_inc(x_141); +x_142 = lean_ctor_get(x_138, 2); +lean_inc(x_142); +x_143 = lean_nat_dec_lt(x_140, x_141); +if (x_143 == 0) { -lean_object* x_17; lean_object* x_18; lean_object* x_19; +lean_object* x_144; lean_object* x_145; +lean_dec(x_142); +lean_dec(x_141); +lean_dec(x_140); lean_dec(x_13); -lean_dec(x_11); -lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_17 = lean_ctor_get(x_1, 0); -lean_inc(x_17); -lean_dec(x_1); -x_18 = lean_ctor_get(x_17, 1); -lean_inc(x_18); -lean_dec(x_17); -x_19 = lean_apply_2(x_18, lean_box(0), x_15); -return x_19; +x_144 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_144, 0, x_138); +lean_ctor_set(x_144, 1, x_139); +x_145 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_145, 0, x_144); +lean_ctor_set(x_145, 1, x_10); +return x_145; } else { -lean_object* x_20; uint8_t x_21; -x_20 = lean_unsigned_to_nat(0u); -x_21 = lean_nat_dec_eq(x_12, x_20); -if (x_21 == 0) +lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; +if (lean_is_exclusive(x_138)) { + lean_ctor_release(x_138, 0); + lean_ctor_release(x_138, 1); + lean_ctor_release(x_138, 2); + x_146 = x_138; +} else { + lean_dec_ref(x_138); + x_146 = lean_box(0); +} +x_147 = lean_nat_add(x_140, x_142); +if (lean_is_scalar(x_146)) { + x_148 = lean_alloc_ctor(0, 3, 0); +} else { + x_148 = x_146; +} +lean_ctor_set(x_148, 0, x_147); +lean_ctor_set(x_148, 1, x_141); +lean_ctor_set(x_148, 2, x_142); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_149 = lean_infer_type(x_13, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_149) == 0) { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; size_t x_27; size_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; -lean_dec(x_15); -x_22 = lean_unsigned_to_nat(1u); -x_23 = lean_nat_sub(x_12, x_22); -x_24 = lean_ctor_get(x_1, 1); -lean_inc(x_24); -x_25 = lean_array_fget(x_2, x_13); -x_26 = lean_array_get_size(x_3); -x_27 = lean_usize_of_nat(x_26); -lean_dec(x_26); -x_28 = 0; -x_29 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2___rarg___closed__2; -lean_inc(x_3); -lean_inc(x_25); -lean_inc(x_5); -lean_inc(x_1); -x_30 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg(x_1, x_5, x_25, x_3, x_27, x_28, x_29); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_13); -lean_inc(x_2); +lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; uint8_t x_158; uint8_t x_159; uint8_t x_160; uint8_t x_161; uint8_t x_162; uint8_t x_163; uint8_t x_164; uint8_t x_165; uint8_t x_166; uint8_t x_167; uint8_t x_168; lean_object* x_169; uint8_t x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; +x_150 = lean_ctor_get(x_6, 0); +lean_inc(x_150); +x_151 = lean_ctor_get(x_149, 0); +lean_inc(x_151); +x_152 = lean_ctor_get(x_149, 1); +lean_inc(x_152); +lean_dec(x_149); +x_153 = lean_ctor_get(x_6, 1); +lean_inc(x_153); +x_154 = lean_ctor_get(x_6, 2); +lean_inc(x_154); +x_155 = lean_ctor_get(x_6, 3); +lean_inc(x_155); +x_156 = lean_ctor_get(x_6, 4); +lean_inc(x_156); +x_157 = lean_ctor_get(x_6, 5); +lean_inc(x_157); +x_158 = lean_ctor_get_uint8(x_150, 0); +x_159 = lean_ctor_get_uint8(x_150, 1); +x_160 = lean_ctor_get_uint8(x_150, 2); +x_161 = lean_ctor_get_uint8(x_150, 3); +x_162 = lean_ctor_get_uint8(x_150, 4); +x_163 = lean_ctor_get_uint8(x_150, 5); +x_164 = lean_ctor_get_uint8(x_150, 6); +x_165 = lean_ctor_get_uint8(x_150, 7); +x_166 = lean_ctor_get_uint8(x_150, 8); +x_167 = lean_ctor_get_uint8(x_150, 10); +x_168 = lean_ctor_get_uint8(x_150, 11); +if (lean_is_exclusive(x_150)) { + x_169 = x_150; +} else { + lean_dec_ref(x_150); + x_169 = lean_box(0); +} +x_170 = 2; +if (lean_is_scalar(x_169)) { + x_171 = lean_alloc_ctor(0, 0, 12); +} else { + x_171 = x_169; +} +lean_ctor_set_uint8(x_171, 0, x_158); +lean_ctor_set_uint8(x_171, 1, x_159); +lean_ctor_set_uint8(x_171, 2, x_160); +lean_ctor_set_uint8(x_171, 3, x_161); +lean_ctor_set_uint8(x_171, 4, x_162); +lean_ctor_set_uint8(x_171, 5, x_163); +lean_ctor_set_uint8(x_171, 6, x_164); +lean_ctor_set_uint8(x_171, 7, x_165); +lean_ctor_set_uint8(x_171, 8, x_166); +lean_ctor_set_uint8(x_171, 9, x_170); +lean_ctor_set_uint8(x_171, 10, x_167); +lean_ctor_set_uint8(x_171, 11, x_168); +x_172 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_172, 0, x_171); +lean_ctor_set(x_172, 1, x_153); +lean_ctor_set(x_172, 2, x_154); +lean_ctor_set(x_172, 3, x_155); +lean_ctor_set(x_172, 4, x_156); +lean_ctor_set(x_172, 5, x_157); +x_173 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getNatParams___spec__1___closed__1; +lean_inc(x_9); lean_inc(x_8); -lean_inc(x_1); -x_31 = lean_alloc_closure((void*)(l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2___rarg___lambda__3), 8, 7); -lean_closure_set(x_31, 0, x_1); -lean_closure_set(x_31, 1, x_8); -lean_closure_set(x_31, 2, x_2); -lean_closure_set(x_31, 3, x_13); -lean_closure_set(x_31, 4, x_4); -lean_closure_set(x_31, 5, x_25); -lean_closure_set(x_31, 6, x_5); -lean_inc(x_5); -x_32 = lean_apply_4(x_5, lean_box(0), lean_box(0), x_30, x_31); -x_33 = lean_alloc_closure((void*)(l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2___rarg___lambda__4), 14, 13); -lean_closure_set(x_33, 0, x_1); -lean_closure_set(x_33, 1, x_13); -lean_closure_set(x_33, 2, x_11); -lean_closure_set(x_33, 3, x_2); -lean_closure_set(x_33, 4, x_3); -lean_closure_set(x_33, 5, x_4); -lean_closure_set(x_33, 6, x_5); -lean_closure_set(x_33, 7, x_6); -lean_closure_set(x_33, 8, x_7); -lean_closure_set(x_33, 9, x_8); -lean_closure_set(x_33, 10, x_9); -lean_closure_set(x_33, 11, x_10); -lean_closure_set(x_33, 12, x_23); -x_34 = lean_apply_4(x_24, lean_box(0), lean_box(0), x_32, x_33); -return x_34; +lean_inc(x_7); +x_174 = l_Lean_Meta_isExprDefEq(x_151, x_173, x_172, x_7, x_8, x_9, x_152); +if (lean_obj_tag(x_174) == 0) +{ +lean_object* x_175; uint8_t x_176; +x_175 = lean_ctor_get(x_174, 0); +lean_inc(x_175); +x_176 = lean_unbox(x_175); +lean_dec(x_175); +if (x_176 == 0) +{ +lean_object* x_177; lean_object* x_178; size_t x_179; size_t x_180; +lean_dec(x_140); +x_177 = lean_ctor_get(x_174, 1); +lean_inc(x_177); +lean_dec(x_174); +x_178 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_178, 0, x_148); +lean_ctor_set(x_178, 1, x_139); +x_179 = 1; +x_180 = lean_usize_add(x_4, x_179); +x_4 = x_180; +x_5 = x_178; +x_10 = x_177; +goto _start; } else { -lean_object* x_35; lean_object* x_36; lean_object* x_37; -lean_dec(x_13); -lean_dec(x_11); -lean_dec(x_10); +lean_object* x_182; lean_object* x_183; lean_object* x_184; size_t x_185; size_t x_186; +x_182 = lean_ctor_get(x_174, 1); +lean_inc(x_182); +lean_dec(x_174); +x_183 = lean_array_push(x_139, x_140); +x_184 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_184, 0, x_148); +lean_ctor_set(x_184, 1, x_183); +x_185 = 1; +x_186 = lean_usize_add(x_4, x_185); +x_4 = x_186; +x_5 = x_184; +x_10 = x_182; +goto _start; +} +} +else +{ +lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; +lean_dec(x_148); +lean_dec(x_140); +lean_dec(x_139); 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_35 = lean_ctor_get(x_1, 0); -lean_inc(x_35); -lean_dec(x_1); -x_36 = lean_ctor_get(x_35, 1); -lean_inc(x_36); -lean_dec(x_35); -x_37 = lean_apply_2(x_36, lean_box(0), x_15); -return x_37; +x_188 = lean_ctor_get(x_174, 0); +lean_inc(x_188); +x_189 = lean_ctor_get(x_174, 1); +lean_inc(x_189); +if (lean_is_exclusive(x_174)) { + lean_ctor_release(x_174, 0); + lean_ctor_release(x_174, 1); + x_190 = x_174; +} else { + lean_dec_ref(x_174); + x_190 = lean_box(0); } +if (lean_is_scalar(x_190)) { + x_191 = lean_alloc_ctor(1, 2, 0); +} else { + x_191 = x_190; } +lean_ctor_set(x_191, 0, x_188); +lean_ctor_set(x_191, 1, x_189); +return x_191; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2(lean_object* x_1, lean_object* x_2) { -_start: +else { -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2___rarg___boxed), 15, 0); -return x_3; +lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; +lean_dec(x_148); +lean_dec(x_140); +lean_dec(x_139); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +x_192 = lean_ctor_get(x_149, 0); +lean_inc(x_192); +x_193 = lean_ctor_get(x_149, 1); +lean_inc(x_193); +if (lean_is_exclusive(x_149)) { + lean_ctor_release(x_149, 0); + lean_ctor_release(x_149, 1); + x_194 = x_149; +} else { + lean_dec_ref(x_149); + x_194 = lean_box(0); } +if (lean_is_scalar(x_194)) { + x_195 = lean_alloc_ctor(1, 2, 0); +} else { + x_195 = x_194; } -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve_go___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_4 = lean_ctor_get(x_1, 0); -lean_inc(x_4); -lean_dec(x_1); -x_5 = lean_ctor_get(x_4, 1); -lean_inc(x_5); -lean_dec(x_4); -x_6 = lean_apply_2(x_5, lean_box(0), x_2); -return x_6; +lean_ctor_set(x_195, 0, x_192); +lean_ctor_set(x_195, 1, x_193); +return x_195; } } -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve_go___rarg___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_getNatParams___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_5; -x_5 = lean_ctor_get(x_4, 0); -lean_inc(x_5); -lean_dec(x_4); -if (lean_obj_tag(x_5) == 0) +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; size_t x_19; size_t x_20; lean_object* x_21; +x_10 = lean_array_get_size(x_3); +x_11 = l_Array_toSubarray___rarg(x_3, x_1, x_10); +x_12 = l_Array_ofSubarray___rarg(x_11); +x_13 = lean_array_get_size(x_12); +x_14 = lean_unsigned_to_nat(0u); +x_15 = lean_unsigned_to_nat(1u); +lean_inc(x_13); +x_16 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_13); +lean_ctor_set(x_16, 2, x_15); +x_17 = l_Lean_Elab_WF_GuessLex_naryVarNames___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_usize_of_nat(x_13); +lean_dec(x_13); +x_20 = 0; +x_21 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getNatParams___spec__1(x_2, x_12, x_19, x_20, x_18, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_12); +if (lean_obj_tag(x_21) == 0) { -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -lean_inc(x_1); -x_6 = lean_alloc_closure((void*)(l_Lean_Elab_WF_GuessLex_solve_go___rarg___lambda__1___boxed), 3, 2); -lean_closure_set(x_6, 0, x_1); -lean_closure_set(x_6, 1, x_2); -x_7 = lean_ctor_get(x_1, 0); -lean_inc(x_7); -lean_dec(x_1); -x_8 = lean_ctor_get(x_7, 1); -lean_inc(x_8); -lean_dec(x_7); -x_9 = lean_box(0); -x_10 = lean_apply_2(x_8, lean_box(0), x_9); -x_11 = lean_apply_4(x_3, lean_box(0), lean_box(0), x_10, x_6); -return x_11; -} -else +uint8_t x_22; +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -lean_dec(x_3); -lean_dec(x_2); -x_12 = lean_ctor_get(x_5, 0); -lean_inc(x_12); -lean_dec(x_5); -x_13 = lean_ctor_get(x_1, 0); -lean_inc(x_13); -lean_dec(x_1); -x_14 = lean_ctor_get(x_13, 1); -lean_inc(x_14); -lean_dec(x_13); -x_15 = lean_apply_2(x_14, lean_box(0), x_12); -return x_15; -} -} +lean_object* x_23; lean_object* x_24; +x_23 = lean_ctor_get(x_21, 0); +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +lean_dec(x_23); +lean_ctor_set(x_21, 0, x_24); +return x_21; } -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve_go___rarg___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +else { -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_6 = lean_ctor_get(x_1, 1); -lean_inc(x_6); -x_7 = lean_array_get_size(x_2); -x_8 = lean_unsigned_to_nat(0u); -x_9 = lean_unsigned_to_nat(1u); -lean_inc(x_7); -x_10 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_10, 0, x_8); -lean_ctor_set(x_10, 1, x_7); -lean_ctor_set(x_10, 2, x_9); -x_11 = lean_box(0); -x_12 = l_Lean_Elab_WF_GuessLex_evalRecCall___lambda__2___closed__4; -lean_inc_n(x_7, 2); -lean_inc(x_6); -lean_inc(x_1); -x_13 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2___rarg(x_1, x_2, x_3, x_4, x_6, x_7, x_10, x_12, x_8, x_7, x_9, x_7, x_8, lean_box(0), x_12); -lean_dec(x_7); -lean_inc(x_6); -x_14 = lean_alloc_closure((void*)(l_Lean_Elab_WF_GuessLex_solve_go___rarg___lambda__2), 4, 3); -lean_closure_set(x_14, 0, x_1); -lean_closure_set(x_14, 1, x_11); -lean_closure_set(x_14, 2, x_6); -x_15 = lean_apply_4(x_6, lean_box(0), lean_box(0), x_13, x_14); -return x_15; +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_25 = lean_ctor_get(x_21, 0); +x_26 = lean_ctor_get(x_21, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_21); +x_27 = lean_ctor_get(x_25, 1); +lean_inc(x_27); +lean_dec(x_25); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_26); +return x_28; } } -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve_go___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: +else { -lean_object* x_5; uint8_t x_6; -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_1); -x_5 = lean_alloc_closure((void*)(l_Lean_Elab_WF_GuessLex_solve_go___rarg___lambda__3___boxed), 5, 4); -lean_closure_set(x_5, 0, x_1); -lean_closure_set(x_5, 1, x_2); -lean_closure_set(x_5, 2, x_3); -lean_closure_set(x_5, 3, x_4); -x_6 = l_Array_isEmpty___rarg(x_3); -lean_dec(x_3); -if (x_6 == 0) +uint8_t x_29; +x_29 = !lean_is_exclusive(x_21); +if (x_29 == 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_dec(x_4); -x_7 = lean_ctor_get(x_1, 1); -lean_inc(x_7); -x_8 = lean_ctor_get(x_1, 0); -lean_inc(x_8); -lean_dec(x_1); -x_9 = lean_ctor_get(x_8, 1); -lean_inc(x_9); -lean_dec(x_8); -x_10 = lean_box(0); -x_11 = lean_apply_2(x_9, lean_box(0), x_10); -x_12 = lean_apply_4(x_7, lean_box(0), lean_box(0), x_11, x_5); -return x_12; +return x_21; } else { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -lean_dec(x_5); -x_13 = lean_ctor_get(x_1, 0); -lean_inc(x_13); -lean_dec(x_1); -x_14 = lean_ctor_get(x_13, 1); -lean_inc(x_14); -lean_dec(x_13); -x_15 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_15, 0, x_4); -x_16 = lean_apply_2(x_14, lean_box(0), x_15); -return x_16; -} +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_21, 0); +x_31 = lean_ctor_get(x_21, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_21); +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_Elab_WF_GuessLex_solve_go(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Lean_Elab_WF_GuessLex_solve_go___rarg), 4, 0); -return x_3; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +static lean_object* _init_l_Lean_Elab_WF_GuessLex_getNatParams___closed__1() { _start: { -lean_object* x_6; -x_6 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg___lambda__1(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_5); -return x_6; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Core_instMonadCoreM; +x_2 = l_ReaderT_instMonadReaderT___rarg(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l_Lean_Elab_WF_GuessLex_getNatParams___closed__2() { _start: { -lean_object* x_5; -x_5 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg___lambda__2(x_1, x_2, x_3, x_4); -lean_dec(x_4); -return x_5; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_WF_GuessLex_getNatParams___closed__1; +x_2 = l_ReaderT_instApplicativeReaderT___rarg(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___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) { +static lean_object* _init_l_Lean_Elab_WF_GuessLex_getNatParams___closed__3() { _start: { -uint8_t x_8; lean_object* x_9; -x_8 = lean_unbox(x_7); -lean_dec(x_7); -x_9 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_8); -return x_9; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_WF_GuessLex_getNatParams___closed__2; +x_2 = lean_ctor_get(x_1, 1); +lean_inc(x_2); +x_3 = l_instMonadControlT__1___rarg(x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___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_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_getNatParams(lean_object* x_1, 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; -x_8 = lean_unbox_usize(x_2); +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; +x_8 = lean_ctor_get(x_2, 5); +lean_inc(x_8); lean_dec(x_2); -x_9 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_10 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg___lambda__4(x_1, x_8, x_3, x_4, x_5, x_9, x_7); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -size_t x_8; size_t x_9; lean_object* x_10; -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_10 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg(x_1, x_2, x_3, x_4, x_8, x_9, x_7); -return x_10; +x_9 = l_Lean_Elab_WF_GuessLex_getNatParams___closed__3; +x_10 = lean_alloc_closure((void*)(l_Lean_Elab_WF_GuessLex_getNatParams___lambda__1___boxed), 9, 2); +lean_closure_set(x_10, 0, x_1); +lean_closure_set(x_10, 1, x_9); +x_11 = 0; +x_12 = l_Lean_Meta_lambdaTelescope___at_Lean_PrettyPrinter_Delaborator_returnsPi___spec__1___rarg(x_8, x_10, x_11, x_3, x_4, x_5, x_6, x_7); +return x_12; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2___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_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getNatParams___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) { _start: { -lean_object* x_9; -x_9 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2___rarg___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_8); +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_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getNatParams___spec__1(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_9; +lean_dec(x_1); +return x_13; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___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, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_getNatParams___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_16; -x_16 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2___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, x_15); -lean_dec(x_12); -return x_16; +lean_object* x_10; +x_10 = l_Lean_Elab_WF_GuessLex_getNatParams___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_2); +return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve_go___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT uint8_t l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_isForbidden(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_4; -x_4 = l_Lean_Elab_WF_GuessLex_solve_go___rarg___lambda__1(x_1, x_2, x_3); -lean_dec(x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve_go___rarg___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_4; uint8_t x_5; +x_4 = lean_array_get_size(x_1); +x_5 = lean_nat_dec_lt(x_2, x_4); +lean_dec(x_4); +if (x_5 == 0) { -lean_object* x_6; -x_6 = l_Lean_Elab_WF_GuessLex_solve_go___rarg___lambda__3(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_5); +uint8_t x_6; +x_6 = 0; return x_6; } -} -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -lean_object* x_4; lean_object* x_5; -x_4 = l_Lean_Elab_WF_GuessLex_naryVarNames___closed__1; -x_5 = l_Lean_Elab_WF_GuessLex_solve_go___rarg(x_1, x_2, x_3, x_4); -return x_5; -} +lean_object* x_7; uint8_t x_8; +x_7 = lean_array_fget(x_1, x_2); +x_8 = l_Array_contains___at___private_Lean_Meta_FunInfo_0__Lean_Meta_collectDeps_visit___spec__2(x_7, x_3); +lean_dec(x_7); +return x_8; } -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Lean_Elab_WF_GuessLex_solve___rarg), 3, 0); -return x_3; } } -static lean_object* _init_l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__1() { +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_isForbidden___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Parser", 6); -return x_1; +uint8_t x_4; lean_object* x_5; +x_4 = l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_isForbidden(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_5 = lean_box(x_4); +return x_5; } } -static lean_object* _init_l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__2() { +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform___spec__1(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Term", 4); -return x_1; +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_3, x_4); +if (x_5 == 0) +{ +lean_object* x_6; uint8_t x_7; +x_6 = lean_array_uget(x_2, x_3); +x_7 = l_Array_contains___at___private_Lean_Meta_FunInfo_0__Lean_Meta_collectDeps_visit___spec__2(x_6, x_1); +lean_dec(x_6); +if (x_7 == 0) +{ +size_t x_8; size_t x_9; +x_8 = 1; +x_9 = lean_usize_add(x_3, x_8); +x_3 = x_9; +goto _start; } +else +{ +uint8_t x_11; +x_11 = 1; +return x_11; } -static lean_object* _init_l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__3() { -_start: +} +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("tuple", 5); -return x_1; +uint8_t x_12; +x_12 = 0; +return x_12; } } -static lean_object* _init_l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__4() { +} +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform___spec__2(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_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_WF_GuessLex_initFn____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_9____closed__6; -x_2 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__1; -x_3 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__2; -x_4 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__3; -x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); -return x_5; +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_3, x_4); +if (x_5 == 0) +{ +lean_object* x_6; uint8_t x_7; +x_6 = lean_array_uget(x_2, x_3); +x_7 = lean_nat_dec_lt(x_1, x_6); +lean_dec(x_6); +if (x_7 == 0) +{ +uint8_t x_8; +x_8 = 1; +return x_8; } +else +{ +size_t x_9; size_t x_10; +x_9 = 1; +x_10 = lean_usize_add(x_3, x_9); +x_3 = x_10; +goto _start; } -static lean_object* _init_l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__5() { -_start: +} +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("(", 1); -return x_1; +uint8_t x_12; +x_12 = 0; +return x_12; } } -static lean_object* _init_l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__6() { +} +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform___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_1; -x_1 = lean_mk_string_from_bytes("null", 4); -return x_1; +lean_object* x_6; lean_object* x_7; lean_object* x_8; +lean_dec(x_4); +x_6 = lean_unsigned_to_nat(1u); +x_7 = lean_nat_add(x_1, x_6); +lean_dec(x_1); +x_8 = l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform(x_2, x_3, x_7, x_5); +return x_8; } } -static lean_object* _init_l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__7() { +static lean_object* _init_l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform___closed__1() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; +lean_object* x_1; lean_object* x_2; x_1 = lean_box(0); -x_2 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__6; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } } -static lean_object* _init_l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__8() { +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes(",", 1); -return x_1; +lean_object* x_5; lean_object* x_6; lean_object* x_24; uint8_t x_25; +x_5 = lean_array_get_size(x_2); +x_24 = lean_unsigned_to_nat(0u); +x_25 = lean_nat_dec_lt(x_24, x_5); +if (x_25 == 0) +{ +lean_object* x_26; +x_26 = lean_box(0); +x_6 = x_26; +goto block_23; } +else +{ +size_t x_27; size_t x_28; uint8_t x_29; +x_27 = 0; +x_28 = lean_usize_of_nat(x_5); +x_29 = l_Array_anyMUnsafe_any___at_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform___spec__2(x_3, x_2, x_27, x_28); +if (x_29 == 0) +{ +lean_object* x_30; +x_30 = lean_box(0); +x_6 = x_30; +goto block_23; } -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_mkTupleSyntax(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_31; lean_object* x_32; +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_31 = l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform___closed__1; +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_4); +return x_32; +} +} +block_23: { lean_object* x_7; lean_object* x_8; uint8_t x_9; +lean_dec(x_6); x_7 = lean_array_get_size(x_1); x_8 = lean_unsigned_to_nat(0u); -x_9 = lean_nat_dec_eq(x_7, x_8); +x_9 = lean_nat_dec_lt(x_8, x_7); if (x_9 == 0) { -lean_object* x_10; uint8_t x_11; -x_10 = lean_unsigned_to_nat(1u); -x_11 = lean_nat_dec_eq(x_7, x_10); -if (x_11 == 0) -{ -lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_12 = lean_ctor_get(x_4, 5); -lean_inc(x_12); -lean_dec(x_4); -x_13 = 0; -x_14 = l_Lean_SourceInfo_fromRef(x_12, x_13); -x_15 = lean_st_ref_get(x_5, x_6); -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +lean_dec(x_7); +lean_inc(x_3); +x_10 = lean_mk_array(x_5, x_3); +x_11 = lean_array_push(x_4, x_10); +x_12 = lean_box(0); +x_13 = l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform___lambda__1(x_3, x_1, x_2, x_12, x_11); +return x_13; +} +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; lean_object* x_25; lean_object* x_26; 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_ctor_get(x_15, 0); -lean_dec(x_17); -x_18 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__5; -lean_inc(x_14); -x_19 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_19, 0, x_14); -lean_ctor_set(x_19, 1, x_18); -x_20 = lean_nat_dec_lt(x_8, x_7); +size_t x_14; size_t x_15; uint8_t x_16; +x_14 = 0; +x_15 = lean_usize_of_nat(x_7); lean_dec(x_7); -x_21 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__8; -lean_inc(x_14); -x_22 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_22, 0, x_14); -lean_ctor_set(x_22, 1, x_21); -x_23 = lean_array_get_size(x_1); -lean_inc(x_1); -x_24 = l_Array_toSubarray___rarg(x_1, x_10, x_23); -x_25 = l_Array_ofSubarray___rarg(x_24); -x_26 = l_Lean_Syntax_SepArray_ofElems(x_21, x_25); -lean_dec(x_25); -x_27 = l_Lean_Elab_WF_GuessLex_naryVarNames___closed__1; -x_28 = l_Array_append___rarg(x_27, x_26); -x_29 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__7; -lean_inc(x_14); -x_30 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_30, 0, x_14); -lean_ctor_set(x_30, 1, x_29); -lean_ctor_set(x_30, 2, x_28); -x_31 = l_Lean_Elab_WF_GuessLex_collectRecCalls___lambda__2___closed__7; -lean_inc(x_14); -x_32 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_32, 0, x_14); -lean_ctor_set(x_32, 1, x_31); -if (x_20 == 0) +x_16 = l_Array_anyMUnsafe_any___at_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform___spec__1(x_3, x_1, x_14, x_15); +if (x_16 == 0) { -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -lean_dec(x_1); -x_33 = lean_box(0); -x_34 = l___private_Init_Util_0__outOfBounds___rarg(x_33); -lean_inc(x_14); -x_35 = l_Lean_Syntax_node3(x_14, x_29, x_34, x_22, x_30); -x_36 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__4; -x_37 = l_Lean_Syntax_node3(x_14, x_36, x_19, x_35, x_32); -lean_ctor_set(x_15, 0, x_37); -return x_15; +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +lean_inc(x_3); +x_17 = lean_mk_array(x_5, x_3); +x_18 = lean_array_push(x_4, x_17); +x_19 = lean_box(0); +x_20 = l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform___lambda__1(x_3, x_1, x_2, x_19, x_18); +return x_20; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_38 = lean_array_fget(x_1, x_8); +lean_object* x_21; lean_object* x_22; +lean_dec(x_5); +x_21 = lean_box(0); +x_22 = l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform___lambda__1(x_3, x_1, x_2, x_21, x_4); +return x_22; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform___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; 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_Elab_WF_GuessLex_generateCombinations_x3f_goUniform___spec__1(x_1, x_2, x_5, x_6); +lean_dec(x_2); lean_dec(x_1); -lean_inc(x_14); -x_39 = l_Lean_Syntax_node3(x_14, x_29, x_38, x_22, x_30); -x_40 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__4; -x_41 = l_Lean_Syntax_node3(x_14, x_40, x_19, x_39, x_32); -lean_ctor_set(x_15, 0, x_41); -return x_15; +x_8 = lean_box(x_7); +return x_8; } } -else +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_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; lean_object* x_56; lean_object* x_57; -x_42 = lean_ctor_get(x_15, 1); -lean_inc(x_42); -lean_dec(x_15); -x_43 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__5; -lean_inc(x_14); -x_44 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_44, 0, x_14); -lean_ctor_set(x_44, 1, x_43); -x_45 = lean_nat_dec_lt(x_8, x_7); -lean_dec(x_7); -x_46 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__8; -lean_inc(x_14); -x_47 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_47, 0, x_14); -lean_ctor_set(x_47, 1, x_46); -x_48 = lean_array_get_size(x_1); -lean_inc(x_1); -x_49 = l_Array_toSubarray___rarg(x_1, x_10, x_48); -x_50 = l_Array_ofSubarray___rarg(x_49); -x_51 = l_Lean_Syntax_SepArray_ofElems(x_46, x_50); -lean_dec(x_50); -x_52 = l_Lean_Elab_WF_GuessLex_naryVarNames___closed__1; -x_53 = l_Array_append___rarg(x_52, x_51); -x_54 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__7; -lean_inc(x_14); -x_55 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_55, 0, x_14); -lean_ctor_set(x_55, 1, x_54); -lean_ctor_set(x_55, 2, x_53); -x_56 = l_Lean_Elab_WF_GuessLex_collectRecCalls___lambda__2___closed__7; -lean_inc(x_14); -x_57 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_57, 0, x_14); -lean_ctor_set(x_57, 1, x_56); -if (x_45 == 0) -{ -lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; -lean_dec(x_1); -x_58 = lean_box(0); -x_59 = l___private_Init_Util_0__outOfBounds___rarg(x_58); -lean_inc(x_14); -x_60 = l_Lean_Syntax_node3(x_14, x_54, x_59, x_47, x_55); -x_61 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__4; -x_62 = l_Lean_Syntax_node3(x_14, x_61, x_44, x_60, x_57); -x_63 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_63, 0, x_62); -lean_ctor_set(x_63, 1, x_42); -return x_63; -} -else -{ -lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; -x_64 = lean_array_fget(x_1, x_8); +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_Elab_WF_GuessLex_generateCombinations_x3f_goUniform___spec__2(x_1, x_2, x_5, x_6); +lean_dec(x_2); lean_dec(x_1); -lean_inc(x_14); -x_65 = l_Lean_Syntax_node3(x_14, x_54, x_64, x_47, x_55); -x_66 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__4; -x_67 = l_Lean_Syntax_node3(x_14, x_66, x_44, x_65, x_57); -x_68 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_68, 0, x_67); -lean_ctor_set(x_68, 1, x_42); -return x_68; +x_8 = lean_box(x_7); +return x_8; } } +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_WF_GuessLex_generateCombinations_x3f_go___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5) { +_start: +{ +uint8_t x_6; +x_6 = lean_usize_dec_eq(x_4, x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_7 = lean_array_uget(x_3, x_4); +x_8 = lean_unsigned_to_nat(0u); +x_9 = lean_nat_dec_lt(x_8, x_2); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_10 = l_instInhabitedNat; +x_11 = l___private_Init_Util_0__outOfBounds___rarg(x_10); +x_12 = lean_nat_dec_eq(x_7, x_11); +lean_dec(x_11); +lean_dec(x_7); +if (x_12 == 0) +{ +uint8_t x_13; +x_13 = 1; +return x_13; } else { -lean_object* x_69; lean_object* x_70; -lean_dec(x_7); -lean_dec(x_4); -x_69 = lean_array_fget(x_1, x_8); -lean_dec(x_1); -x_70 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_70, 0, x_69); -lean_ctor_set(x_70, 1, x_6); -return x_70; +size_t x_14; size_t x_15; +x_14 = 1; +x_15 = lean_usize_add(x_4, x_14); +x_4 = x_15; +goto _start; } } else { -lean_object* x_71; uint8_t x_72; lean_object* x_73; lean_object* x_74; uint8_t x_75; +lean_object* x_17; uint8_t x_18; +x_17 = lean_array_fget(x_1, x_8); +x_18 = lean_nat_dec_eq(x_7, x_17); +lean_dec(x_17); lean_dec(x_7); -lean_dec(x_1); -x_71 = lean_ctor_get(x_4, 5); -lean_inc(x_71); -lean_dec(x_4); -x_72 = 0; -x_73 = l_Lean_SourceInfo_fromRef(x_71, x_72); -x_74 = lean_st_ref_get(x_5, x_6); -x_75 = !lean_is_exclusive(x_74); -if (x_75 == 0) +if (x_18 == 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; -x_76 = lean_ctor_get(x_74, 0); -lean_dec(x_76); -x_77 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__5; -lean_inc(x_73); -x_78 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_78, 0, x_73); -lean_ctor_set(x_78, 1, x_77); -x_79 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__7; -x_80 = l_Lean_Elab_WF_GuessLex_naryVarNames___closed__1; -lean_inc(x_73); -x_81 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_81, 0, x_73); -lean_ctor_set(x_81, 1, x_79); -lean_ctor_set(x_81, 2, x_80); -x_82 = l_Lean_Elab_WF_GuessLex_collectRecCalls___lambda__2___closed__7; -lean_inc(x_73); -x_83 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_83, 0, x_73); -lean_ctor_set(x_83, 1, x_82); -x_84 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__4; -x_85 = l_Lean_Syntax_node3(x_73, x_84, x_78, x_81, x_83); -lean_ctor_set(x_74, 0, x_85); -return x_74; +uint8_t x_19; +x_19 = 1; +return x_19; } else { -lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; -x_86 = lean_ctor_get(x_74, 1); -lean_inc(x_86); -lean_dec(x_74); -x_87 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__5; -lean_inc(x_73); -x_88 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_88, 0, x_73); -lean_ctor_set(x_88, 1, x_87); -x_89 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__7; -x_90 = l_Lean_Elab_WF_GuessLex_naryVarNames___closed__1; -lean_inc(x_73); -x_91 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_91, 0, x_73); -lean_ctor_set(x_91, 1, x_89); -lean_ctor_set(x_91, 2, x_90); -x_92 = l_Lean_Elab_WF_GuessLex_collectRecCalls___lambda__2___closed__7; -lean_inc(x_73); -x_93 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_93, 0, x_73); -lean_ctor_set(x_93, 1, x_92); -x_94 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__4; -x_95 = l_Lean_Syntax_node3(x_73, x_94, x_88, x_91, x_93); -x_96 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_96, 0, x_95); -lean_ctor_set(x_96, 1, x_86); -return x_96; -} +size_t x_20; size_t x_21; +x_20 = 1; +x_21 = lean_usize_add(x_4, x_20); +x_4 = x_21; +goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_mkTupleSyntax___boxed(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_Elab_WF_GuessLex_mkTupleSyntax(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -return x_7; +uint8_t x_23; +x_23 = 0; +return x_23; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__1(size_t x_1, size_t x_2, lean_object* x_3) { +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_generateCombinations_x3f_go___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -uint8_t x_4; -x_4 = lean_usize_dec_lt(x_2, x_1); -if (x_4 == 0) +uint8_t x_12; +x_12 = lean_nat_dec_le(x_7, x_6); +if (x_12 == 0) { -return x_3; -} -else +lean_object* x_13; uint8_t x_14; +x_13 = lean_unsigned_to_nat(0u); +x_14 = lean_nat_dec_eq(x_5, x_13); +if (x_14 == 0) { -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; -x_5 = lean_array_uget(x_3, x_2); -x_6 = lean_unsigned_to_nat(0u); -x_7 = lean_array_uset(x_3, x_2, x_6); -x_8 = lean_mk_syntax_ident(x_5); -x_9 = 1; -x_10 = lean_usize_add(x_2, x_9); -x_11 = lean_array_uset(x_7, x_2, x_8); -x_2 = x_10; -x_3 = x_11; -goto _start; -} -} -} -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__4___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: +lean_object* x_15; lean_object* x_16; uint8_t x_17; +lean_dec(x_9); +x_15 = lean_unsigned_to_nat(1u); +x_16 = lean_nat_sub(x_5, x_15); +lean_dec(x_5); +x_17 = l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_isForbidden(x_1, x_4, x_6); +if (x_17 == 0) { -lean_object* x_11; lean_object* x_12; +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_18 = lean_nat_add(x_4, x_15); +lean_inc(x_6); +lean_inc(x_10); +x_19 = lean_array_push(x_10, x_6); lean_inc(x_3); -x_11 = l_Lean_resolveGlobalName___at_Lean_Elab_WF_GuessLex_naryVarNames_freshen___spec__1(x_3, x_6, x_7, x_8, x_9, x_10); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -if (lean_obj_tag(x_12) == 0) -{ -uint8_t x_13; -x_13 = !lean_is_exclusive(x_11); -if (x_13 == 0) +lean_inc(x_2); +lean_inc(x_1); +x_20 = l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_go(x_1, x_2, x_3, x_18, x_19, x_11); +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +if (lean_obj_tag(x_21) == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_14 = lean_ctor_get(x_11, 0); -lean_dec(x_14); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_3); -lean_ctor_set(x_15, 1, x_4); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_1); -lean_ctor_set(x_16, 1, x_15); -x_17 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set(x_11, 0, x_17); -return x_11; +uint8_t x_22; +lean_dec(x_16); +lean_dec(x_10); +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_22 = !lean_is_exclusive(x_20); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; +x_23 = lean_ctor_get(x_20, 0); +lean_dec(x_23); +x_24 = lean_box(0); +lean_ctor_set(x_20, 0, x_24); +return x_20; } 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_11, 1); -lean_inc(x_18); -lean_dec(x_11); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_3); -lean_ctor_set(x_19, 1, x_4); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_1); -lean_ctor_set(x_20, 1, x_19); -x_21 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_21, 0, 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_18); -return x_22; +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_20, 1); +lean_inc(x_25); +lean_dec(x_20); +x_26 = lean_box(0); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_25); +return x_27; } } else { -lean_object* x_23; lean_object* x_24; -x_23 = lean_ctor_get(x_12, 0); -lean_inc(x_23); -x_24 = lean_ctor_get(x_12, 1); -lean_inc(x_24); -lean_dec(x_12); -if (lean_obj_tag(x_24) == 0) -{ -uint8_t x_25; -x_25 = !lean_is_exclusive(x_11); -if (x_25 == 0) -{ -lean_object* x_26; lean_object* x_27; uint8_t x_28; -x_26 = lean_ctor_get(x_11, 0); -lean_dec(x_26); -x_27 = lean_ctor_get(x_23, 0); -lean_inc(x_27); -lean_dec(x_23); -x_28 = lean_name_eq(x_27, x_2); -lean_dec(x_27); -if (x_28 == 0) -{ -lean_object* x_29; lean_object* x_30; lean_object* x_31; -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_1); -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_11, 0, x_31); -return x_11; +lean_object* x_28; lean_object* x_29; lean_object* x_30; +lean_dec(x_21); +x_28 = lean_ctor_get(x_20, 1); +lean_inc(x_28); +lean_dec(x_20); +x_29 = lean_nat_add(x_6, x_8); +lean_dec(x_6); +x_30 = lean_box(0); +x_5 = x_16; +x_6 = x_29; +x_9 = x_30; +x_11 = x_28; +goto _start; +} } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -lean_dec(x_1); -lean_inc(x_3); -x_32 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_32, 0, x_3); -x_33 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_33, 0, x_32); -x_34 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_34, 0, x_3); -lean_ctor_set(x_34, 1, x_4); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -x_36 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_11, 0, x_36); -return x_11; +lean_object* x_32; lean_object* x_33; +x_32 = lean_nat_add(x_6, x_8); +lean_dec(x_6); +x_33 = lean_box(0); +x_5 = x_16; +x_6 = x_32; +x_9 = x_33; +goto _start; } } else { -lean_object* x_37; lean_object* x_38; uint8_t x_39; -x_37 = lean_ctor_get(x_11, 1); -lean_inc(x_37); -lean_dec(x_11); -x_38 = lean_ctor_get(x_23, 0); -lean_inc(x_38); -lean_dec(x_23); -x_39 = lean_name_eq(x_38, x_2); -lean_dec(x_38); -if (x_39 == 0) -{ -lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_3); -lean_ctor_set(x_40, 1, x_4); -x_41 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_41, 0, x_1); -lean_ctor_set(x_41, 1, x_40); -x_42 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_42, 0, x_41); -x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_42); -lean_ctor_set(x_43, 1, x_37); -return x_43; +lean_object* x_35; lean_object* x_36; +lean_dec(x_10); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_35 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_35, 0, x_9); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_11); +return x_36; +} } 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_37; lean_object* x_38; +lean_dec(x_10); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); -lean_inc(x_3); -x_44 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_44, 0, x_3); -x_45 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_45, 0, x_44); -x_46 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_46, 0, x_3); -lean_ctor_set(x_46, 1, x_4); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_45); -lean_ctor_set(x_47, 1, x_46); -x_48 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_48, 0, 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_37); -return x_49; +x_37 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_37, 0, x_9); +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_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_go___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -uint8_t x_50; -lean_dec(x_24); -lean_dec(x_23); -x_50 = !lean_is_exclusive(x_11); -if (x_50 == 0) +lean_object* x_5; uint8_t x_6; +x_5 = lean_array_get_size(x_4); +x_6 = lean_nat_dec_lt(x_1, x_5); +lean_dec(x_5); +if (x_6 == 0) { -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_51 = lean_ctor_get(x_11, 0); -lean_dec(x_51); -x_52 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_52, 0, x_3); -lean_ctor_set(x_52, 1, x_4); -x_53 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_53, 0, x_1); -lean_ctor_set(x_53, 1, x_52); -x_54 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_11, 0, x_54); -return x_11; +lean_object* x_7; lean_object* x_8; +x_7 = l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform___closed__1; +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_4); +return x_8; } else { -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_55 = lean_ctor_get(x_11, 1); -lean_inc(x_55); -lean_dec(x_11); -x_56 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_56, 0, x_3); -lean_ctor_set(x_56, 1, x_4); -x_57 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_57, 0, x_1); -lean_ctor_set(x_57, 1, x_56); -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_55); -return x_59; -} -} +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_4); +return x_10; } } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__4___closed__1() { +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_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) { _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_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -uint8_t x_13; -x_13 = lean_nat_dec_le(x_5, x_4); -if (x_13 == 0) -{ -lean_object* x_14; uint8_t x_15; -x_14 = lean_unsigned_to_nat(0u); -x_15 = lean_nat_dec_eq(x_3, x_14); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_16 = lean_unsigned_to_nat(1u); -x_17 = lean_nat_sub(x_3, x_16); -lean_dec(x_3); -x_18 = lean_ctor_get(x_7, 1); -lean_inc(x_18); +lean_object* x_7; uint8_t x_8; +x_7 = lean_array_get_size(x_2); +x_8 = lean_nat_dec_lt(x_4, x_7); lean_dec(x_7); -x_19 = lean_ctor_get(x_18, 1); -lean_inc(x_19); -if (lean_obj_tag(x_19) == 0) +if (x_8 == 0) { -uint8_t x_20; -lean_dec(x_17); -lean_dec(x_10); +lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_dec(x_4); lean_dec(x_2); -x_20 = !lean_is_exclusive(x_18); -if (x_20 == 0) +lean_dec(x_1); +x_9 = lean_array_get_size(x_5); +x_10 = lean_unsigned_to_nat(0u); +x_11 = lean_nat_dec_lt(x_10, x_9); +if (x_11 == 0) { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_21 = lean_ctor_get(x_18, 1); -lean_dec(x_21); -x_22 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__4___closed__1; -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_18); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_12); -return x_24; +lean_object* x_12; lean_object* x_13; +lean_dec(x_9); +x_12 = lean_box(0); +x_13 = l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_go___lambda__1(x_3, x_12, x_5, x_6); +lean_dec(x_5); +lean_dec(x_3); +return x_13; } else { -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_18, 0); -lean_inc(x_25); -lean_dec(x_18); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_19); -x_27 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__4___closed__1; -x_28 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_28, 0, x_27); -lean_ctor_set(x_28, 1, x_26); -x_29 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_29, 0, x_28); -lean_ctor_set(x_29, 1, x_12); -return x_29; -} +size_t x_14; size_t x_15; uint8_t x_16; +x_14 = 0; +x_15 = lean_usize_of_nat(x_9); +x_16 = l_Array_anyMUnsafe_any___at_Lean_Elab_WF_GuessLex_generateCombinations_x3f_go___spec__1(x_5, x_9, x_5, x_14, x_15); +lean_dec(x_9); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +x_17 = lean_box(0); +x_18 = l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_go___lambda__1(x_3, x_17, x_5, x_6); +lean_dec(x_5); +lean_dec(x_3); +return x_18; } 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; -x_30 = lean_ctor_get(x_18, 0); -lean_inc(x_30); -lean_dec(x_18); -x_31 = lean_ctor_get(x_19, 0); -lean_inc(x_31); -x_32 = lean_ctor_get(x_19, 1); -lean_inc(x_32); -lean_dec(x_19); -x_33 = l_Lean_Name_append(x_31, x_30); -x_34 = lean_box(0); -lean_inc(x_10); -lean_inc(x_2); -x_35 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__4___lambda__1(x_2, x_1, x_33, x_32, x_34, x_8, x_9, x_10, x_11, x_12); -x_36 = lean_ctor_get(x_35, 0); -lean_inc(x_36); -if (lean_obj_tag(x_36) == 0) +lean_object* x_19; lean_object* x_20; lean_object* x_21; +lean_inc(x_5); +x_19 = lean_array_push(x_6, x_5); +x_20 = lean_box(0); +x_21 = l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_go___lambda__1(x_3, x_20, x_5, x_19); +lean_dec(x_5); +lean_dec(x_3); +return x_21; +} +} +} +else { -uint8_t x_37; -lean_dec(x_17); -lean_dec(x_10); +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_array_fget(x_2, x_4); +x_23 = lean_unsigned_to_nat(0u); +x_24 = lean_unsigned_to_nat(1u); +x_25 = lean_box(0); +lean_inc(x_22); +x_26 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_generateCombinations_x3f_go___spec__2(x_1, x_2, x_3, x_4, x_22, x_23, x_22, x_24, x_25, x_5, x_6); +lean_dec(x_22); lean_dec(x_4); -lean_dec(x_2); -x_37 = !lean_is_exclusive(x_35); -if (x_37 == 0) +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +if (lean_obj_tag(x_27) == 0) { -lean_object* x_38; lean_object* x_39; -x_38 = lean_ctor_get(x_35, 0); -lean_dec(x_38); -x_39 = lean_ctor_get(x_36, 0); -lean_inc(x_39); -lean_dec(x_36); -lean_ctor_set(x_35, 0, x_39); -return x_35; +uint8_t x_28; +x_28 = !lean_is_exclusive(x_26); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; +x_29 = lean_ctor_get(x_26, 0); +lean_dec(x_29); +x_30 = lean_box(0); +lean_ctor_set(x_26, 0, x_30); +return x_26; } else { -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_35, 1); -lean_inc(x_40); -lean_dec(x_35); -x_41 = lean_ctor_get(x_36, 0); -lean_inc(x_41); -lean_dec(x_36); -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_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_26, 1); +lean_inc(x_31); +lean_dec(x_26); +x_32 = lean_box(0); +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 { -lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_43 = lean_ctor_get(x_35, 1); -lean_inc(x_43); +uint8_t x_34; +lean_dec(x_27); +x_34 = !lean_is_exclusive(x_26); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; +x_35 = lean_ctor_get(x_26, 0); lean_dec(x_35); -x_44 = lean_ctor_get(x_36, 0); -lean_inc(x_44); -lean_dec(x_36); -x_45 = lean_nat_add(x_4, x_6); -lean_dec(x_4); -x_3 = x_17; -x_4 = x_45; -x_7 = x_44; -x_12 = x_43; -goto _start; +x_36 = l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform___closed__1; +lean_ctor_set(x_26, 0, x_36); +return x_26; } +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_26, 1); +lean_inc(x_37); +lean_dec(x_26); +x_38 = l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform___closed__1; +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_37); +return x_39; } } -else +} +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_WF_GuessLex_generateCombinations_x3f_go___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: { -lean_object* x_47; -lean_dec(x_10); +size_t x_6; size_t x_7; uint8_t x_8; lean_object* x_9; +x_6 = lean_unbox_usize(x_4); lean_dec(x_4); +x_7 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_8 = l_Array_anyMUnsafe_any___at_Lean_Elab_WF_GuessLex_generateCombinations_x3f_go___spec__1(x_1, x_2, x_3, x_6, x_7); lean_dec(x_3); lean_dec(x_2); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_7); -lean_ctor_set(x_47, 1, x_12); -return x_47; +lean_dec(x_1); +x_9 = lean_box(x_8); +return x_9; } } -else +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_generateCombinations_x3f_go___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) { +_start: { -lean_object* x_48; -lean_dec(x_10); +lean_object* x_12; +x_12 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_generateCombinations_x3f_go___spec__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_8); +lean_dec(x_7); lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_48 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_48, 0, x_7); -lean_ctor_set(x_48, 1, x_12); -return x_48; -} +return x_12; } } -LEAN_EXPORT lean_object* l_Lean_unresolveNameGlobal_unresolveNameCore___at_Lean_Elab_WF_GuessLex_buildTermWF___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_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_go___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 = 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_object* x_5; +x_5 = l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_go___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_Lean_unresolveNameGlobal_unresolveNameCore___at_Lean_Elab_WF_GuessLex_buildTermWF___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_Elab_WF_GuessLex_generateCombinations_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _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; -x_8 = l_Lean_Name_componentsRev(x_2); -x_9 = lean_unsigned_to_nat(0u); -x_10 = l_List_lengthTRAux___rarg(x_8, x_9); -x_11 = lean_box(0); -x_12 = lean_box(0); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_12); -lean_ctor_set(x_13, 1, x_8); -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_unsigned_to_nat(1u); -lean_inc(x_10); -x_16 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__4(x_1, x_11, x_10, x_9, x_10, x_15, x_14, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_10); -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -lean_dec(x_17); -if (lean_obj_tag(x_18) == 0) -{ -uint8_t x_19; -x_19 = !lean_is_exclusive(x_16); -if (x_19 == 0) +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_4 = lean_unsigned_to_nat(0u); +x_5 = l_Lean_Elab_WF_GuessLex_naryVarNames___closed__1; +lean_inc(x_2); +lean_inc(x_1); +x_6 = l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform(x_1, x_2, x_4, x_5); +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +x_8 = !lean_is_exclusive(x_7); +if (x_8 == 0) { -lean_object* x_20; -x_20 = lean_ctor_get(x_16, 0); -lean_dec(x_20); -lean_ctor_set(x_16, 0, x_11); -return x_16; +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_ctor_get(x_7, 0); +lean_dec(x_9); +x_10 = lean_ctor_get(x_6, 1); +lean_inc(x_10); +lean_dec(x_6); +x_11 = l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_go(x_1, x_2, x_3, x_4, x_5, x_10); +x_12 = lean_ctor_get(x_11, 1); +lean_inc(x_12); +lean_dec(x_11); +lean_ctor_set(x_7, 0, x_12); +return x_7; } else { -lean_object* x_21; lean_object* x_22; -x_21 = lean_ctor_get(x_16, 1); -lean_inc(x_21); -lean_dec(x_16); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_11); -lean_ctor_set(x_22, 1, x_21); -return x_22; +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_dec(x_7); +x_13 = lean_ctor_get(x_6, 1); +lean_inc(x_13); +lean_dec(x_6); +x_14 = l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_go(x_1, x_2, x_3, x_4, x_5, x_13); +x_15 = lean_ctor_get(x_14, 1); +lean_inc(x_15); +lean_dec(x_14); +x_16 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_16, 0, x_15); +return x_16; } } -else +} +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_WF_GuessLex_generateMeasures___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_23; -x_23 = !lean_is_exclusive(x_16); -if (x_23 == 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_24; lean_object* x_25; -x_24 = lean_ctor_get(x_16, 0); -lean_dec(x_24); -x_25 = lean_ctor_get(x_18, 0); -lean_inc(x_25); -lean_dec(x_18); -lean_ctor_set(x_16, 0, x_25); -return x_16; +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_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_16, 1); -lean_inc(x_26); -lean_dec(x_16); -x_27 = lean_ctor_get(x_18, 0); -lean_inc(x_27); -lean_dec(x_18); -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; -} +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; } } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_buildTermWF___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_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_generateMeasures___spec__2(size_t x_1, size_t x_2, lean_object* x_3) { _start: { -uint8_t x_12; -x_12 = lean_usize_dec_lt(x_5, x_4); -if (x_12 == 0) +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) { -lean_object* x_13; -lean_dec(x_9); -lean_dec(x_2); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_6); -lean_ctor_set(x_13, 1, x_11); -return x_13; +return x_3; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; -lean_dec(x_6); -x_14 = lean_array_uget(x_3, x_5); -lean_inc(x_9); -x_15 = l_Lean_unresolveNameGlobal_unresolveNameCore___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__3(x_1, x_14, x_7, x_8, x_9, x_10, x_11); -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; size_t x_18; size_t x_19; -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_5, x_18); -lean_inc(x_2); -{ -size_t _tmp_4 = x_19; -lean_object* _tmp_5 = x_2; -lean_object* _tmp_10 = x_17; -x_5 = _tmp_4; -x_6 = _tmp_5; -x_11 = _tmp_10; -} +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_8, 0, x_5); +x_9 = 1; +x_10 = lean_usize_add(x_2, x_9); +x_11 = lean_array_uset(x_7, x_2, x_8); +x_2 = x_10; +x_3 = x_11; goto _start; } -else -{ -uint8_t x_21; -lean_dec(x_9); -lean_dec(x_2); -x_21 = !lean_is_exclusive(x_15); -if (x_21 == 0) +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_generateMeasures___spec__3(size_t x_1, size_t x_2, lean_object* x_3) { +_start: { -lean_object* x_22; uint8_t x_23; -x_22 = lean_ctor_get(x_15, 0); -lean_dec(x_22); -x_23 = !lean_is_exclusive(x_16); -if (x_23 == 0) +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) { -lean_object* x_24; lean_object* x_25; -x_24 = lean_box(0); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_16); -lean_ctor_set(x_25, 1, x_24); -lean_ctor_set(x_15, 0, x_25); -return x_15; +return x_3; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_26 = lean_ctor_get(x_16, 0); -lean_inc(x_26); -lean_dec(x_16); -x_27 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_27, 0, x_26); -x_28 = lean_box(0); -x_29 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_28); -lean_ctor_set(x_15, 0, x_29); -return x_15; +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_8, 0, x_5); +x_9 = 1; +x_10 = lean_usize_add(x_2, x_9); +x_11 = lean_array_uset(x_7, x_2, x_8); +x_2 = x_10; +x_3 = x_11; +goto _start; } } -else +} +static lean_object* _init_l_Lean_Elab_WF_GuessLex_generateMeasures___closed__1() { +_start: { -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_30 = lean_ctor_get(x_15, 1); -lean_inc(x_30); -lean_dec(x_15); -x_31 = lean_ctor_get(x_16, 0); -lean_inc(x_31); -if (lean_is_exclusive(x_16)) { - lean_ctor_release(x_16, 0); - x_32 = x_16; -} else { - lean_dec_ref(x_16); - x_32 = lean_box(0); +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Too many combinations", 21); +return x_1; } -if (lean_is_scalar(x_32)) { - x_33 = lean_alloc_ctor(1, 1, 0); -} else { - x_33 = x_32; } -lean_ctor_set(x_33, 0, x_31); -x_34 = lean_box(0); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_30); -return x_36; +static lean_object* _init_l_Lean_Elab_WF_GuessLex_generateMeasures___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_WF_GuessLex_generateMeasures___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } +static lean_object* _init_l_Lean_Elab_WF_GuessLex_generateMeasures___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_WF_GuessLex_naryVarNames___closed__1; +x_2 = lean_array_get_size(x_1); +return x_2; +} } +static size_t _init_l_Lean_Elab_WF_GuessLex_generateMeasures___closed__4() { +_start: +{ +lean_object* x_1; size_t x_2; +x_1 = l_Lean_Elab_WF_GuessLex_generateMeasures___closed__3; +x_2 = lean_usize_of_nat(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_unresolveNameGlobal___at_Lean_Elab_WF_GuessLex_buildTermWF___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_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_generateMeasures(lean_object* x_1, 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; size_t x_20; size_t x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; -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_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; size_t x_15; size_t x_16; lean_object* x_17; +x_8 = lean_unsigned_to_nat(32u); +lean_inc(x_2); +x_9 = l_Lean_Elab_WF_GuessLex_generateCombinations_x3f(x_1, x_2, x_8); +x_10 = lean_ctor_get(x_9, 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); -lean_inc(x_1); -x_12 = l_Lean_getRevAliases(x_11, x_1); -x_13 = l_List_redLength___rarg(x_12); -x_14 = lean_mk_empty_array_with_capacity(x_13); -lean_dec(x_13); -x_15 = l_List_toArrayAux___rarg(x_12, x_14); -x_16 = l_Lean_rootNamespace; -lean_inc(x_1); -x_17 = l_Lean_Name_append(x_16, x_1); -x_18 = lean_array_push(x_15, 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_Lean_Elab_WF_GuessLex_evalRecCall___lambda__2___closed__4; -x_23 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__5(x_1, x_22, x_18, x_20, x_21, x_22, x_3, x_4, x_5, x_6, x_10); -lean_dec(x_18); -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_24, 0); -lean_inc(x_25); -lean_dec(x_24); -if (lean_obj_tag(x_25) == 0) -{ -uint8_t x_26; -x_26 = !lean_is_exclusive(x_23); -if (x_26 == 0) +x_11 = lean_array_get_size(x_2); +lean_dec(x_2); +x_12 = lean_unsigned_to_nat(1u); +x_13 = lean_nat_dec_lt(x_12, x_11); +x_14 = lean_array_get_size(x_10); +x_15 = lean_usize_of_nat(x_14); +lean_dec(x_14); +x_16 = 0; +x_17 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_generateMeasures___spec__2(x_15, x_16, x_10); +if (x_13 == 0) { -lean_object* x_27; -x_27 = lean_ctor_get(x_23, 0); -lean_dec(x_27); -lean_ctor_set(x_23, 0, x_1); -return x_23; +size_t x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +lean_dec(x_11); +x_18 = l_Lean_Elab_WF_GuessLex_generateMeasures___closed__4; +x_19 = l_Lean_Elab_WF_GuessLex_naryVarNames___closed__1; +x_20 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_generateMeasures___spec__3(x_18, x_16, x_19); +x_21 = l_Array_append___rarg(x_17, 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_7); +return x_22; } else { -lean_object* x_28; lean_object* x_29; -x_28 = lean_ctor_get(x_23, 1); -lean_inc(x_28); -lean_dec(x_23); -x_29 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_29, 0, x_1); -lean_ctor_set(x_29, 1, x_28); -return x_29; +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; size_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_23 = l_List_range(x_11); +x_24 = l_List_redLength___rarg(x_23); +x_25 = lean_mk_empty_array_with_capacity(x_24); +lean_dec(x_24); +x_26 = l_List_toArrayAux___rarg(x_23, x_25); +x_27 = lean_array_get_size(x_26); +x_28 = lean_usize_of_nat(x_27); +lean_dec(x_27); +x_29 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_generateMeasures___spec__3(x_28, x_16, x_26); +x_30 = l_Array_append___rarg(x_17, 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_7); +return x_31; } } -else -{ -uint8_t x_30; -lean_dec(x_1); -x_30 = !lean_is_exclusive(x_23); -if (x_30 == 0) -{ -lean_object* x_31; lean_object* x_32; -x_31 = lean_ctor_get(x_23, 0); -lean_dec(x_31); -x_32 = lean_ctor_get(x_25, 0); -lean_inc(x_32); -lean_dec(x_25); -lean_ctor_set(x_23, 0, x_32); -return x_23; } -else +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_WF_GuessLex_generateMeasures___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_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_23, 1); -lean_inc(x_33); -lean_dec(x_23); -x_34 = lean_ctor_get(x_25, 0); -lean_inc(x_34); -lean_dec(x_25); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_33); -return x_35; +lean_object* x_7; +x_7 = l_Lean_throwError___at_Lean_Elab_WF_GuessLex_generateMeasures___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_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_generateMeasures___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_generateMeasures___spec__2(x_4, x_5, x_3); +return x_6; } } -LEAN_EXPORT lean_object* l_Lean_unresolveNameGlobal___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__2___lambda__2(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_generateMeasures___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_dec(x_3); -if (x_1 == 0) -{ -lean_object* x_9; lean_object* x_10; -x_9 = lean_box(0); -x_10 = l_Lean_unresolveNameGlobal___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__2___lambda__1(x_2, x_9, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_4); -return x_10; +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_generateMeasures___spec__3(x_4, x_5, x_3); +return x_6; } -else +} +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_generateMeasures___boxed(lean_object* x_1, 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_11; lean_object* x_12; -lean_inc(x_2); -x_11 = l_Lean_resolveGlobalName___at_Lean_Elab_WF_GuessLex_naryVarNames_freshen___spec__1(x_2, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_7); +lean_object* x_8; +x_8 = l_Lean_Elab_WF_GuessLex_generateMeasures(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); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -if (lean_obj_tag(x_12) == 0) -{ -uint8_t x_13; -x_13 = !lean_is_exclusive(x_11); -if (x_13 == 0) +lean_dec(x_3); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg___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_14; -x_14 = lean_ctor_get(x_11, 0); -lean_dec(x_14); -lean_ctor_set(x_11, 0, x_2); -return x_11; +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_6 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_6, 0, x_1); +lean_ctor_set(x_6, 1, x_2); +x_7 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_7, 0, x_3); +lean_ctor_set(x_7, 1, x_6); +x_8 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_8, 0, x_7); +x_9 = lean_apply_2(x_4, lean_box(0), x_8); +return x_9; } -else +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg___lambda__2(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; -x_15 = lean_ctor_get(x_11, 1); -lean_inc(x_15); -lean_dec(x_11); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_2); -lean_ctor_set(x_16, 1, x_15); -return x_16; +uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_5 = 1; +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_1); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_2); +lean_ctor_set(x_8, 1, x_7); +x_9 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_9, 0, x_8); +x_10 = lean_apply_2(x_3, lean_box(0), x_9); +return x_10; } } -else +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___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) { +_start: { -lean_object* x_17; lean_object* x_18; -x_17 = lean_ctor_get(x_12, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_12, 1); -lean_inc(x_18); -lean_dec(x_12); -if (lean_obj_tag(x_18) == 0) -{ -uint8_t x_19; -x_19 = !lean_is_exclusive(x_11); -if (x_19 == 0) -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_11, 0); -lean_dec(x_20); -x_21 = lean_ctor_get(x_17, 0); -lean_inc(x_21); -lean_dec(x_17); -lean_inc(x_21); -x_22 = lean_private_to_user_name(x_21); -if (lean_obj_tag(x_22) == 0) -{ -uint8_t x_23; -x_23 = lean_name_eq(x_21, x_2); -lean_dec(x_21); -if (x_23 == 0) -{ -lean_object* x_24; lean_object* x_25; -x_24 = l_Lean_rootNamespace; -x_25 = l_Lean_Name_append(x_24, x_2); -lean_ctor_set(x_11, 0, x_25); -return x_11; -} -else +uint8_t x_8; uint8_t x_9; +x_8 = 0; +x_9 = l_Lean_Elab_WF_GuessLex_instDecidableEqGuessLexRel(x_7, x_8); +if (x_9 == 0) { -lean_ctor_set(x_11, 0, x_2); -return x_11; -} -} -else +lean_object* x_10; uint8_t x_11; uint8_t x_12; +x_10 = lean_array_push(x_1, x_2); +x_11 = 2; +x_12 = l_Lean_Elab_WF_GuessLex_instDecidableEqGuessLexRel(x_7, x_11); +if (x_12 == 0) { -lean_object* x_26; uint8_t x_27; -lean_dec(x_21); -x_26 = lean_ctor_get(x_22, 0); -lean_inc(x_26); -lean_dec(x_22); -x_27 = lean_name_eq(x_26, x_2); -lean_dec(x_26); -if (x_27 == 0) +uint8_t x_13; uint8_t x_14; +x_13 = 1; +x_14 = l_Lean_Elab_WF_GuessLex_instDecidableEqGuessLexRel(x_7, x_13); +if (x_14 == 0) { -lean_object* x_28; lean_object* x_29; -x_28 = l_Lean_rootNamespace; -x_29 = l_Lean_Name_append(x_28, x_2); -lean_ctor_set(x_11, 0, x_29); -return x_11; +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; +lean_dec(x_6); +lean_dec(x_5); +x_15 = lean_ctor_get(x_3, 0); +lean_inc(x_15); +lean_dec(x_3); +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec(x_15); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_4); +lean_ctor_set(x_17, 1, x_10); +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_17); +x_21 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_21, 0, x_20); +x_22 = lean_apply_2(x_16, lean_box(0), x_21); +return x_22; } else { -lean_ctor_set(x_11, 0, x_2); -return x_11; -} +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_3, 0); +lean_inc(x_23); +lean_dec(x_3); +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +lean_dec(x_23); +x_25 = lean_box(0); +lean_inc(x_24); +x_26 = lean_apply_2(x_24, lean_box(0), x_25); +x_27 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg___lambda__1___boxed), 5, 4); +lean_closure_set(x_27, 0, x_4); +lean_closure_set(x_27, 1, x_10); +lean_closure_set(x_27, 2, x_5); +lean_closure_set(x_27, 3, x_24); +x_28 = lean_apply_4(x_6, lean_box(0), lean_box(0), x_26, x_27); +return x_28; } } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_11, 1); +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_29 = lean_ctor_get(x_3, 0); +lean_inc(x_29); +lean_dec(x_3); +x_30 = lean_ctor_get(x_29, 1); lean_inc(x_30); -lean_dec(x_11); -x_31 = lean_ctor_get(x_17, 0); -lean_inc(x_31); -lean_dec(x_17); -lean_inc(x_31); -x_32 = lean_private_to_user_name(x_31); -if (lean_obj_tag(x_32) == 0) -{ -uint8_t x_33; -x_33 = lean_name_eq(x_31, x_2); -lean_dec(x_31); -if (x_33 == 0) -{ -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = l_Lean_rootNamespace; -x_35 = l_Lean_Name_append(x_34, x_2); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_30); -return x_36; -} -else -{ -lean_object* x_37; -x_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_2); -lean_ctor_set(x_37, 1, x_30); -return x_37; -} +lean_dec(x_29); +x_31 = lean_box(0); +lean_inc(x_30); +x_32 = lean_apply_2(x_30, lean_box(0), x_31); +x_33 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg___lambda__1___boxed), 5, 4); +lean_closure_set(x_33, 0, x_4); +lean_closure_set(x_33, 1, x_10); +lean_closure_set(x_33, 2, x_5); +lean_closure_set(x_33, 3, x_30); +x_34 = lean_apply_4(x_6, lean_box(0), lean_box(0), x_32, x_33); +return x_34; } -else -{ -lean_object* x_38; uint8_t x_39; -lean_dec(x_31); -x_38 = lean_ctor_get(x_32, 0); -lean_inc(x_38); -lean_dec(x_32); -x_39 = lean_name_eq(x_38, x_2); -lean_dec(x_38); -if (x_39 == 0) -{ -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = l_Lean_rootNamespace; -x_41 = l_Lean_Name_append(x_40, x_2); -x_42 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_30); -return x_42; } else { -lean_object* x_43; -x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_2); -lean_ctor_set(x_43, 1, x_30); -return x_43; -} +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +lean_dec(x_4); +lean_dec(x_2); +x_35 = lean_ctor_get(x_3, 0); +lean_inc(x_35); +lean_dec(x_3); +x_36 = lean_ctor_get(x_35, 1); +lean_inc(x_36); +lean_dec(x_35); +x_37 = lean_box(0); +lean_inc(x_36); +x_38 = lean_apply_2(x_36, lean_box(0), x_37); +x_39 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg___lambda__2___boxed), 4, 3); +lean_closure_set(x_39, 0, x_1); +lean_closure_set(x_39, 1, x_5); +lean_closure_set(x_39, 2, x_36); +x_40 = lean_apply_4(x_6, lean_box(0), lean_box(0), x_38, x_39); +return x_40; } } } -else +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg___lambda__4(lean_object* x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, size_t x_6, lean_object* x_7) { +_start: { -uint8_t x_44; -lean_dec(x_18); -lean_dec(x_17); -x_44 = !lean_is_exclusive(x_11); -if (x_44 == 0) +if (lean_obj_tag(x_7) == 0) { -lean_object* x_45; -x_45 = lean_ctor_get(x_11, 0); -lean_dec(x_45); -lean_ctor_set(x_11, 0, x_2); +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +lean_dec(x_7); +x_9 = lean_ctor_get(x_1, 0); +lean_inc(x_9); +lean_dec(x_1); +x_10 = lean_ctor_get(x_9, 1); +lean_inc(x_10); +lean_dec(x_9); +x_11 = lean_apply_2(x_10, lean_box(0), x_8); return x_11; } else { -lean_object* x_46; lean_object* x_47; -x_46 = lean_ctor_get(x_11, 1); -lean_inc(x_46); -lean_dec(x_11); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_2); -lean_ctor_set(x_47, 1, x_46); -return x_47; -} -} -} +lean_object* x_12; size_t x_13; size_t x_14; lean_object* x_15; +x_12 = lean_ctor_get(x_7, 0); +lean_inc(x_12); +lean_dec(x_7); +x_13 = 1; +x_14 = lean_usize_add(x_2, x_13); +x_15 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg(x_1, x_3, x_4, x_5, x_6, x_14, x_12); +return x_15; } } } -LEAN_EXPORT lean_object* l_Lean_unresolveNameGlobal___at_Lean_Elab_WF_GuessLex_buildTermWF___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_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7) { _start: { uint8_t x_8; -x_8 = l_Lean_Name_hasMacroScopes(x_1); +x_8 = lean_usize_dec_lt(x_6, x_5); if (x_8 == 0) { -lean_object* x_9; lean_object* x_10; -x_9 = lean_box(0); -x_10 = l_Lean_unresolveNameGlobal___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__2___lambda__2(x_2, x_1, x_9, x_3, x_4, x_5, x_6, x_7); -return x_10; +lean_object* x_9; lean_object* x_10; lean_object* x_11; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_9 = lean_ctor_get(x_1, 0); +lean_inc(x_9); +lean_dec(x_1); +x_10 = lean_ctor_get(x_9, 1); +lean_inc(x_10); +lean_dec(x_9); +x_11 = lean_apply_2(x_10, lean_box(0), x_7); +return x_11; } else { -lean_object* x_11; -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_11 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_11, 0, x_1); -lean_ctor_set(x_11, 1, x_7); -return 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; +x_12 = lean_array_uget(x_4, x_6); +x_13 = lean_ctor_get(x_7, 1); +lean_inc(x_13); +x_14 = lean_ctor_get(x_1, 1); +lean_inc(x_14); +x_15 = lean_ctor_get(x_7, 0); +lean_inc(x_15); +lean_dec(x_7); +x_16 = lean_ctor_get(x_13, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_13, 1); +lean_inc(x_17); +lean_dec(x_13); +x_18 = lean_array_uget(x_4, x_6); +lean_inc(x_3); +x_19 = lean_apply_1(x_18, x_3); +lean_inc(x_2); +lean_inc(x_1); +x_20 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg___lambda__3___boxed), 7, 6); +lean_closure_set(x_20, 0, x_17); +lean_closure_set(x_20, 1, x_12); +lean_closure_set(x_20, 2, x_1); +lean_closure_set(x_20, 3, x_16); +lean_closure_set(x_20, 4, x_15); +lean_closure_set(x_20, 5, x_2); +lean_inc(x_2); +x_21 = lean_apply_4(x_2, lean_box(0), lean_box(0), x_19, x_20); +x_22 = lean_box_usize(x_6); +x_23 = lean_box_usize(x_5); +x_24 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg___lambda__4___boxed), 7, 6); +lean_closure_set(x_24, 0, x_1); +lean_closure_set(x_24, 1, x_22); +lean_closure_set(x_24, 2, x_2); +lean_closure_set(x_24, 3, x_3); +lean_closure_set(x_24, 4, x_4); +lean_closure_set(x_24, 5, x_23); +x_25 = lean_apply_4(x_14, lean_box(0), lean_box(0), x_21, x_24); +return x_25; } } } -static lean_object* _init_l_Array_anyMUnsafe_any___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__6___closed__1() { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Elab_WF_GuessLex_mkSizeOf___closed__3; -x_3 = l_Lean_Name_str___override(x_1, x_2); +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg___boxed), 7, 0); return x_3; } } -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__6(lean_object* x_1, size_t x_2, size_t x_3) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2___rarg___lambda__1(lean_object* x_1, lean_object* x_2) { _start: { -uint8_t x_4; -x_4 = lean_usize_dec_eq(x_2, x_3); -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_Elab_WF_GuessLex_buildTermWF___spec__6___closed__1; -x_7 = lean_name_eq(x_5, x_6); -lean_dec(x_5); -if (x_7 == 0) -{ -size_t x_8; size_t x_9; -x_8 = 1; -x_9 = lean_usize_add(x_2, x_8); -x_2 = x_9; -goto _start; -} -else -{ -uint8_t x_11; -x_11 = 1; -return x_11; -} -} -else -{ -uint8_t x_12; -x_12 = 0; -return x_12; -} +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_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 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_5, 0, x_2); +x_6 = lean_box(0); +x_7 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_7, 0, x_5); +lean_ctor_set(x_7, 1, x_6); +x_8 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_8, 0, x_7); +x_9 = lean_apply_2(x_4, lean_box(0), x_8); +return x_9; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7___closed__1() { +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2___rarg___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("app", 3); -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 = l_Array_eraseIdx___rarg(x_1, x_2); +x_10 = lean_array_push(x_3, x_4); +lean_inc(x_5); +x_11 = l_Lean_Elab_WF_GuessLex_solve_go___rarg(x_5, x_9, x_6, x_10); +x_12 = lean_alloc_closure((void*)(l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2___rarg___lambda__1), 2, 1); +lean_closure_set(x_12, 0, x_5); +x_13 = lean_apply_4(x_7, lean_box(0), lean_box(0), x_11, x_12); +return x_13; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7___closed__2() { +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2___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) { _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_WF_GuessLex_initFn____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_9____closed__6; -x_2 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__1; -x_3 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__2; -x_4 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7___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_Lean_Elab_WF_GuessLex_buildTermWF___spec__7___closed__3() { -_start: +lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_9 = lean_ctor_get(x_8, 1); +lean_inc(x_9); +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_1; lean_object* x_2; -x_1 = l_Lean_Elab_WF_GuessLex_mkSizeOf___closed__4; -x_2 = lean_mk_syntax_ident(x_1); -return x_2; -} +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +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_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(1, 1, 0); +lean_ctor_set(x_14, 0, x_2); +x_15 = lean_apply_2(x_13, lean_box(0), x_14); +return x_15; } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7___closed__4() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("num", 3); -return x_1; -} -} -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7___closed__5() { -_start: +lean_object* x_16; uint8_t x_17; +x_16 = lean_ctor_get(x_8, 0); +lean_inc(x_16); +lean_dec(x_8); +x_17 = lean_unbox(x_16); +lean_dec(x_16); +if (x_17 == 0) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7___closed__4; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_18 = lean_ctor_get(x_1, 0); +lean_inc(x_18); +lean_dec(x_1); +x_19 = lean_ctor_get(x_18, 1); +lean_inc(x_19); +lean_dec(x_18); +x_20 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_20, 0, x_2); +x_21 = lean_apply_2(x_19, lean_box(0), x_20); +return x_21; } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7___closed__6() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("0", 1); -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_dec(x_2); +x_22 = lean_ctor_get(x_9, 1); +lean_inc(x_22); +lean_dec(x_9); +lean_inc(x_7); +lean_inc(x_1); +x_23 = lean_alloc_closure((void*)(l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2___rarg___lambda__2___boxed), 8, 7); +lean_closure_set(x_23, 0, x_3); +lean_closure_set(x_23, 1, x_4); +lean_closure_set(x_23, 2, x_5); +lean_closure_set(x_23, 3, x_6); +lean_closure_set(x_23, 4, x_1); +lean_closure_set(x_23, 5, x_22); +lean_closure_set(x_23, 6, x_7); +x_24 = lean_ctor_get(x_1, 0); +lean_inc(x_24); +lean_dec(x_1); +x_25 = lean_ctor_get(x_24, 1); +lean_inc(x_25); +lean_dec(x_24); +x_26 = lean_box(0); +x_27 = lean_apply_2(x_25, lean_box(0), x_26); +x_28 = lean_apply_4(x_7, lean_box(0), lean_box(0), x_27, x_23); +return x_28; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7___closed__7() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("1", 1); -return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, lean_object* x_5, size_t x_6, size_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2___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) { _start: { -uint8_t x_14; -x_14 = lean_usize_dec_lt(x_7, x_6); -if (x_14 == 0) +if (lean_obj_tag(x_14) == 0) { -lean_object* x_15; +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_8); -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; -x_16 = lean_array_uget(x_8, x_7); -x_17 = lean_unsigned_to_nat(0u); -x_18 = lean_array_uset(x_8, x_7, x_17); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_26; lean_object* x_27; uint8_t x_28; lean_object* x_29; lean_object* x_30; -x_26 = lean_ctor_get(x_16, 0); -lean_inc(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); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +lean_dec(x_14); +x_16 = lean_ctor_get(x_1, 0); +lean_inc(x_16); +lean_dec(x_1); +x_17 = lean_ctor_get(x_16, 1); +lean_inc(x_17); lean_dec(x_16); -x_27 = lean_array_get_size(x_26); -x_28 = lean_nat_dec_lt(x_3, x_27); -lean_dec(x_27); -x_29 = lean_array_get_size(x_5); -if (x_28 == 0) -{ -lean_object* x_67; lean_object* x_68; uint8_t x_69; -lean_dec(x_26); -x_67 = l_instInhabitedNat; -x_68 = l___private_Init_Util_0__outOfBounds___rarg(x_67); -x_69 = lean_nat_dec_lt(x_68, x_29); -lean_dec(x_29); -if (x_69 == 0) -{ -lean_object* x_70; lean_object* x_71; -lean_dec(x_68); -x_70 = lean_box(0); -x_71 = l___private_Init_Util_0__outOfBounds___rarg(x_70); -x_30 = x_71; -goto block_66; +x_18 = lean_apply_2(x_17, lean_box(0), x_15); +return x_18; } else { -lean_object* x_72; -x_72 = lean_array_fget(x_5, x_68); -lean_dec(x_68); -x_30 = x_72; -goto block_66; +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_14, 0); +lean_inc(x_19); +lean_dec(x_14); +x_20 = lean_nat_add(x_2, x_3); +lean_dec(x_2); +x_21 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2___rarg(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_3, x_13, x_20, lean_box(0), x_19); +lean_dec(x_13); +return x_21; } } -else -{ -lean_object* x_73; uint8_t x_74; -x_73 = lean_array_fget(x_26, x_3); -lean_dec(x_26); -x_74 = lean_nat_dec_lt(x_73, x_29); -lean_dec(x_29); -if (x_74 == 0) -{ -lean_object* x_75; lean_object* x_76; -lean_dec(x_73); -x_75 = lean_box(0); -x_76 = l___private_Init_Util_0__outOfBounds___rarg(x_75); -x_30 = x_76; -goto block_66; } -else +static lean_object* _init_l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2___rarg___closed__1() { +_start: { -lean_object* x_77; -x_77 = lean_array_fget(x_5, x_73); -lean_dec(x_73); -x_30 = x_77; -goto block_66; -} +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l_Lean_Elab_WF_GuessLex_naryVarNames___closed__1; +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; } -block_66: -{ -lean_object* x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_49; -x_31 = l_Lean_Elab_WF_GuessLex_mkSizeOf___closed__4; -x_32 = 0; -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -x_33 = l_Lean_unresolveNameGlobal___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__2(x_31, x_32, x_9, x_10, x_11, x_12, x_13); -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_array_get_size(x_1); -x_37 = lean_nat_dec_lt(x_3, x_36); -lean_dec(x_36); -x_38 = lean_ctor_get(x_11, 5); -lean_inc(x_38); -x_39 = l_Lean_SourceInfo_fromRef(x_38, x_32); -x_40 = lean_st_ref_get(x_12, x_35); -if (x_37 == 0) -{ -lean_object* x_63; lean_object* x_64; -x_63 = l_Lean_Elab_WF_GuessLex_naryVarNames___closed__1; -x_64 = l___private_Init_Util_0__outOfBounds___rarg(x_63); -x_49 = x_64; -goto block_62; } -else +static lean_object* _init_l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2___rarg___closed__2() { +_start: { -lean_object* x_65; -x_65 = lean_array_fget(x_1, x_3); -x_49 = x_65; -goto block_62; +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2___rarg___closed__1; +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; } -block_48: -{ -lean_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_41); -x_42 = lean_ctor_get(x_40, 1); -lean_inc(x_42); -lean_dec(x_40); -x_43 = lean_mk_syntax_ident(x_34); -x_44 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__7; -lean_inc(x_39); -x_45 = l_Lean_Syntax_node1(x_39, x_44, x_30); -x_46 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7___closed__2; -x_47 = l_Lean_Syntax_node2(x_39, x_46, x_43, x_45); -x_19 = x_47; -x_20 = x_42; -goto block_25; } -block_62: +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___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, lean_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_50; uint8_t x_51; -x_50 = lean_array_get_size(x_49); -x_51 = lean_nat_dec_lt(x_17, x_50); -if (x_51 == 0) +uint8_t x_16; +x_16 = lean_nat_dec_lt(x_13, x_10); +if (x_16 == 0) { -lean_object* x_52; -lean_dec(x_50); -lean_dec(x_49); -x_52 = lean_box(0); -x_41 = x_52; -goto block_48; +lean_object* x_17; lean_object* x_18; lean_object* x_19; +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_17 = lean_ctor_get(x_1, 0); +lean_inc(x_17); +lean_dec(x_1); +x_18 = lean_ctor_get(x_17, 1); +lean_inc(x_18); +lean_dec(x_17); +x_19 = lean_apply_2(x_18, lean_box(0), x_15); +return x_19; } else { -size_t x_53; uint8_t x_54; -x_53 = lean_usize_of_nat(x_50); -lean_dec(x_50); -x_54 = l_Array_anyMUnsafe_any___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__6(x_49, x_4, x_53); -lean_dec(x_49); -if (x_54 == 0) +lean_object* x_20; uint8_t x_21; +x_20 = lean_unsigned_to_nat(0u); +x_21 = lean_nat_dec_eq(x_12, x_20); +if (x_21 == 0) { -lean_object* x_55; -x_55 = lean_box(0); -x_41 = x_55; -goto block_48; +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; size_t x_27; size_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +lean_dec(x_15); +x_22 = lean_unsigned_to_nat(1u); +x_23 = lean_nat_sub(x_12, x_22); +x_24 = lean_ctor_get(x_1, 1); +lean_inc(x_24); +x_25 = lean_array_fget(x_2, x_13); +x_26 = lean_array_get_size(x_3); +x_27 = lean_usize_of_nat(x_26); +lean_dec(x_26); +x_28 = 0; +x_29 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2___rarg___closed__2; +lean_inc(x_3); +lean_inc(x_25); +lean_inc(x_5); +lean_inc(x_1); +x_30 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg(x_1, x_5, x_25, x_3, x_27, x_28, x_29); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_13); +lean_inc(x_2); +lean_inc(x_8); +lean_inc(x_1); +x_31 = lean_alloc_closure((void*)(l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2___rarg___lambda__3), 8, 7); +lean_closure_set(x_31, 0, x_1); +lean_closure_set(x_31, 1, x_8); +lean_closure_set(x_31, 2, x_2); +lean_closure_set(x_31, 3, x_13); +lean_closure_set(x_31, 4, x_4); +lean_closure_set(x_31, 5, x_25); +lean_closure_set(x_31, 6, x_5); +lean_inc(x_5); +x_32 = lean_apply_4(x_5, lean_box(0), lean_box(0), x_30, x_31); +x_33 = lean_alloc_closure((void*)(l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2___rarg___lambda__4), 14, 13); +lean_closure_set(x_33, 0, x_1); +lean_closure_set(x_33, 1, x_13); +lean_closure_set(x_33, 2, x_11); +lean_closure_set(x_33, 3, x_2); +lean_closure_set(x_33, 4, x_3); +lean_closure_set(x_33, 5, x_4); +lean_closure_set(x_33, 6, x_5); +lean_closure_set(x_33, 7, x_6); +lean_closure_set(x_33, 8, x_7); +lean_closure_set(x_33, 9, x_8); +lean_closure_set(x_33, 10, x_9); +lean_closure_set(x_33, 11, x_10); +lean_closure_set(x_33, 12, x_23); +x_34 = lean_apply_4(x_24, lean_box(0), lean_box(0), x_32, x_33); +return x_34; } 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_dec(x_34); -x_56 = lean_ctor_get(x_40, 1); -lean_inc(x_56); -lean_dec(x_40); -x_57 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__7; -lean_inc(x_39); -x_58 = l_Lean_Syntax_node1(x_39, x_57, x_30); -x_59 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7___closed__2; -x_60 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7___closed__3; -x_61 = l_Lean_Syntax_node2(x_39, x_59, x_60, x_58); -x_19 = x_61; -x_20 = x_56; -goto block_25; -} -} +lean_object* x_35; lean_object* x_36; lean_object* x_37; +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_35 = lean_ctor_get(x_1, 0); +lean_inc(x_35); +lean_dec(x_1); +x_36 = lean_ctor_get(x_35, 1); +lean_inc(x_36); +lean_dec(x_35); +x_37 = lean_apply_2(x_36, lean_box(0), x_15); +return x_37; } } } -else -{ -lean_object* x_78; uint8_t x_79; -x_78 = lean_ctor_get(x_16, 0); -lean_inc(x_78); -lean_dec(x_16); -x_79 = lean_nat_dec_eq(x_78, x_3); -lean_dec(x_78); -if (x_79 == 0) -{ -lean_object* x_80; uint8_t x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; -x_80 = lean_ctor_get(x_11, 5); -lean_inc(x_80); -x_81 = 0; -x_82 = l_Lean_SourceInfo_fromRef(x_80, x_81); -x_83 = lean_st_ref_get(x_12, x_13); -x_84 = lean_ctor_get(x_83, 1); -lean_inc(x_84); -lean_dec(x_83); -x_85 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7___closed__6; -lean_inc(x_82); -x_86 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_86, 0, x_82); -lean_ctor_set(x_86, 1, x_85); -x_87 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7___closed__5; -x_88 = l_Lean_Syntax_node1(x_82, x_87, x_86); -x_19 = x_88; -x_20 = x_84; -goto block_25; } -else +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_89; uint8_t 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_11, 5); -lean_inc(x_89); -x_90 = 0; -x_91 = l_Lean_SourceInfo_fromRef(x_89, x_90); -x_92 = lean_st_ref_get(x_12, x_13); -x_93 = lean_ctor_get(x_92, 1); -lean_inc(x_93); -lean_dec(x_92); -x_94 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7___closed__7; -lean_inc(x_91); -x_95 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_95, 0, x_91); -lean_ctor_set(x_95, 1, x_94); -x_96 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7___closed__5; -x_97 = l_Lean_Syntax_node1(x_91, x_96, x_95); -x_19 = x_97; -x_20 = x_93; -goto block_25; +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2___rarg___boxed), 15, 0); +return x_3; } } -block_25: +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve_go___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -size_t x_21; size_t x_22; lean_object* x_23; -x_21 = 1; -x_22 = lean_usize_add(x_7, x_21); -x_23 = lean_array_uset(x_18, x_7, x_19); -x_7 = x_22; -x_8 = x_23; -x_13 = x_20; -goto _start; -} -} +lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +lean_dec(x_1); +x_5 = lean_ctor_get(x_4, 1); +lean_inc(x_5); +lean_dec(x_4); +x_6 = lean_apply_2(x_5, lean_box(0), x_2); +return x_6; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__8(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve_go___rarg___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -uint8_t x_4; -x_4 = lean_usize_dec_lt(x_2, x_1); -if (x_4 == 0) -{ -return x_3; -} +lean_object* x_5; +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +lean_dec(x_4); +if (lean_obj_tag(x_5) == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +lean_inc(x_1); +x_6 = lean_alloc_closure((void*)(l_Lean_Elab_WF_GuessLex_solve_go___rarg___lambda__1___boxed), 3, 2); +lean_closure_set(x_6, 0, x_1); +lean_closure_set(x_6, 1, x_2); +x_7 = lean_ctor_get(x_1, 0); +lean_inc(x_7); +lean_dec(x_1); +x_8 = lean_ctor_get(x_7, 1); +lean_inc(x_8); +lean_dec(x_7); +x_9 = lean_box(0); +x_10 = lean_apply_2(x_8, lean_box(0), x_9); +x_11 = lean_apply_4(x_3, lean_box(0), lean_box(0), x_10, x_6); +return x_11; +} else { -lean_object* x_5; lean_object* x_6; lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; -x_5 = lean_array_uget(x_3, x_2); -x_6 = lean_unsigned_to_nat(0u); -x_7 = lean_array_uset(x_3, x_2, x_6); -x_8 = 1; -x_9 = lean_usize_add(x_2, x_8); -x_10 = lean_array_uset(x_7, x_2, x_5); -x_2 = x_9; -x_3 = x_10; -goto _start; +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +lean_dec(x_3); +lean_dec(x_2); +x_12 = lean_ctor_get(x_5, 0); +lean_inc(x_12); +lean_dec(x_5); +x_13 = lean_ctor_get(x_1, 0); +lean_inc(x_13); +lean_dec(x_1); +x_14 = lean_ctor_get(x_13, 1); +lean_inc(x_14); +lean_dec(x_13); +x_15 = lean_apply_2(x_14, lean_box(0), x_12); +return x_15; } } } -LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve_go___rarg___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_14; uint8_t x_15; -x_14 = lean_unsigned_to_nat(0u); -x_15 = lean_nat_dec_eq(x_5, x_14); -if (x_15 == 0) +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); +x_7 = lean_array_get_size(x_2); +x_8 = lean_unsigned_to_nat(0u); +x_9 = lean_unsigned_to_nat(1u); +lean_inc(x_7); +x_10 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_10, 0, x_8); +lean_ctor_set(x_10, 1, x_7); +lean_ctor_set(x_10, 2, x_9); +x_11 = lean_box(0); +x_12 = l_Lean_Elab_WF_GuessLex_evalRecCall___lambda__2___closed__4; +lean_inc_n(x_7, 2); +lean_inc(x_6); +lean_inc(x_1); +x_13 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2___rarg(x_1, x_2, x_3, x_4, x_6, x_7, x_10, x_12, x_8, x_7, x_9, x_7, x_8, lean_box(0), x_12); +lean_dec(x_7); +lean_inc(x_6); +x_14 = lean_alloc_closure((void*)(l_Lean_Elab_WF_GuessLex_solve_go___rarg___lambda__2), 4, 3); +lean_closure_set(x_14, 0, x_1); +lean_closure_set(x_14, 1, x_11); +lean_closure_set(x_14, 2, x_6); +x_15 = lean_apply_4(x_6, lean_box(0), lean_box(0), x_13, x_14); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve_go___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_16; 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; size_t x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; size_t x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_16 = lean_unsigned_to_nat(1u); -x_17 = lean_nat_sub(x_5, x_16); -lean_dec(x_5); -x_18 = lean_array_fget(x_4, x_6); -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_Lean_Elab_WF_GuessLex_buildTermWF___spec__1(x_20, x_21, x_18); -x_23 = lean_array_get_size(x_3); -x_24 = lean_usize_of_nat(x_23); -lean_dec(x_23); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); +lean_object* x_5; uint8_t x_6; +lean_inc(x_4); lean_inc(x_3); -x_25 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7(x_1, x_2, x_6, x_21, x_22, x_24, x_21, x_3, x_9, x_10, x_11, x_12, x_13); -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_11); -x_28 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax(x_26, x_9, x_10, x_11, x_12, x_27); -x_29 = lean_ctor_get(x_28, 0); -lean_inc(x_29); -x_30 = lean_ctor_get(x_28, 1); -lean_inc(x_30); -lean_dec(x_28); -x_31 = lean_array_get_size(x_22); -x_32 = lean_usize_of_nat(x_31); -lean_dec(x_31); -x_33 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__8(x_32, x_21, x_22); -x_34 = lean_box(0); -x_35 = 1; -x_36 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_33); -lean_ctor_set(x_36, 2, x_29); -lean_ctor_set_uint8(x_36, sizeof(void*)*3, x_35); -x_37 = lean_nat_add(x_6, x_16); -lean_dec(x_6); -x_38 = lean_array_push(x_8, x_36); -x_5 = x_17; -x_6 = x_37; -x_7 = lean_box(0); -x_8 = x_38; -x_13 = x_30; -goto _start; +lean_inc(x_1); +x_5 = lean_alloc_closure((void*)(l_Lean_Elab_WF_GuessLex_solve_go___rarg___lambda__3___boxed), 5, 4); +lean_closure_set(x_5, 0, x_1); +lean_closure_set(x_5, 1, x_2); +lean_closure_set(x_5, 2, x_3); +lean_closure_set(x_5, 3, x_4); +x_6 = l_Array_isEmpty___rarg(x_3); +lean_dec(x_3); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +lean_dec(x_4); +x_7 = lean_ctor_get(x_1, 1); +lean_inc(x_7); +x_8 = lean_ctor_get(x_1, 0); +lean_inc(x_8); +lean_dec(x_1); +x_9 = lean_ctor_get(x_8, 1); +lean_inc(x_9); +lean_dec(x_8); +x_10 = lean_box(0); +x_11 = lean_apply_2(x_9, lean_box(0), x_10); +x_12 = lean_apply_4(x_7, lean_box(0), lean_box(0), x_11, x_5); +return x_12; } else { -lean_object* x_40; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_6); +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_dec(x_5); -lean_dec(x_3); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_8); -lean_ctor_set(x_40, 1, x_13); -return x_40; +x_13 = lean_ctor_get(x_1, 0); +lean_inc(x_13); +lean_dec(x_1); +x_14 = lean_ctor_get(x_13, 1); +lean_inc(x_14); +lean_dec(x_13); +x_15 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_15, 0, x_4); +x_16 = lean_apply_2(x_14, lean_box(0), x_15); +return x_16; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_buildTermWF(lean_object* x_1, lean_object* x_2, lean_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_WF_GuessLex_solve_go(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_9 = lean_array_get_size(x_2); -x_10 = lean_mk_empty_array_with_capacity(x_9); -x_11 = lean_unsigned_to_nat(0u); -x_12 = l_Array_mapIdxM_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__9(x_1, x_2, x_3, x_2, x_9, x_11, lean_box(0), x_10, x_4, x_5, x_6, x_7, x_8); -return x_12; +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Lean_Elab_WF_GuessLex_solve_go___rarg), 4, 0); +return x_3; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -size_t x_4; size_t x_5; lean_object* x_6; -x_4 = lean_unbox_usize(x_1); -lean_dec(x_1); -x_5 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__1(x_4, x_5, x_3); +lean_object* x_6; +x_6 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg___lambda__1(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); return x_6; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__4___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_11; -x_11 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__4___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_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_2); -return x_11; +lean_object* x_5; +x_5 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg___lambda__2(x_1, x_2, x_3, x_4); +lean_dec(x_4); +return x_5; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___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) { _start: { -lean_object* x_13; -x_13 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_11); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_1); -return x_13; +uint8_t x_8; lean_object* x_9; +x_8 = lean_unbox(x_7); +lean_dec(x_7); +x_9 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_8); +return x_9; } } -LEAN_EXPORT lean_object* l_Lean_unresolveNameGlobal_unresolveNameCore___at_Lean_Elab_WF_GuessLex_buildTermWF___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_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___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) { _start: { -lean_object* x_8; -x_8 = l_Lean_unresolveNameGlobal_unresolveNameCore___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__3___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); +size_t x_8; size_t x_9; lean_object* x_10; +x_8 = lean_unbox_usize(x_2); lean_dec(x_2); -return x_8; +x_9 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_10 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg___lambda__4(x_1, x_8, x_3, x_4, x_5, x_9, x_7); +return x_10; } } -LEAN_EXPORT lean_object* l_Lean_unresolveNameGlobal_unresolveNameCore___at_Lean_Elab_WF_GuessLex_buildTermWF___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_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_8; -x_8 = l_Lean_unresolveNameGlobal_unresolveNameCore___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +size_t x_8; size_t x_9; lean_object* x_10; +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = lean_unbox_usize(x_6); lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -return x_8; +x_10 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__1___rarg(x_1, x_2, x_3, x_4, x_8, x_9, x_7); +return x_10; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_buildTermWF___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_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2___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) { _start: { -size_t x_12; size_t x_13; lean_object* x_14; -x_12 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_13 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_14 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__5(x_1, x_2, x_3, x_12, x_13, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); +lean_object* x_9; +x_9 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2___rarg___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_3); -lean_dec(x_1); -return x_14; +lean_dec(x_2); +return x_9; } } -LEAN_EXPORT lean_object* l_Lean_unresolveNameGlobal___at_Lean_Elab_WF_GuessLex_buildTermWF___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_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___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, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { -lean_object* x_8; -x_8 = l_Lean_unresolveNameGlobal___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__2___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_8; +lean_object* x_16; +x_16 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2___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, x_15); +lean_dec(x_12); +return x_16; } } -LEAN_EXPORT lean_object* l_Lean_unresolveNameGlobal___at_Lean_Elab_WF_GuessLex_buildTermWF___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_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve_go___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -uint8_t x_9; lean_object* x_10; -x_9 = lean_unbox(x_1); -lean_dec(x_1); -x_10 = l_Lean_unresolveNameGlobal___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__2___lambda__2(x_9, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_10; +lean_object* x_4; +x_4 = l_Lean_Elab_WF_GuessLex_solve_go___rarg___lambda__1(x_1, x_2, x_3); +lean_dec(x_3); +return x_4; } } -LEAN_EXPORT lean_object* l_Lean_unresolveNameGlobal___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve_go___rarg___lambda__3___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_8; lean_object* x_9; -x_8 = lean_unbox(x_2); -lean_dec(x_2); -x_9 = l_Lean_unresolveNameGlobal___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__2(x_1, x_8, x_3, x_4, x_5, x_6, x_7); -return x_9; +lean_object* x_6; +x_6 = l_Lean_Elab_WF_GuessLex_solve_go___rarg___lambda__3(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +return x_6; } } -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve___rarg(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; -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_Elab_WF_GuessLex_buildTermWF___spec__6(x_1, x_4, x_5); -lean_dec(x_1); -x_7 = lean_box(x_6); -return x_7; +lean_object* x_4; lean_object* x_5; +x_4 = l_Lean_Elab_WF_GuessLex_naryVarNames___closed__1; +x_5 = l_Lean_Elab_WF_GuessLex_solve_go___rarg(x_1, x_2, x_3, x_4); +return x_5; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___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_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve(lean_object* x_1, lean_object* x_2) { _start: { -size_t x_14; size_t x_15; size_t x_16; lean_object* x_17; -x_14 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_15 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_16 = lean_unbox_usize(x_7); -lean_dec(x_7); -x_17 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7(x_1, x_2, x_3, x_14, x_5, x_15, x_16, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_17; +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Lean_Elab_WF_GuessLex_solve___rarg), 3, 0); +return x_3; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__1() { _start: { -size_t x_4; size_t x_5; lean_object* x_6; -x_4 = lean_unbox_usize(x_1); -lean_dec(x_1); -x_5 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__8(x_4, x_5, x_3); -return x_6; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Parser", 6); +return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at_Lean_Elab_WF_GuessLex_buildTermWF___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) { +static lean_object* _init_l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__2() { _start: { -lean_object* x_14; -x_14 = l_Array_mapIdxM_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -return x_14; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Term", 4); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_buildTermWF___boxed(lean_object* x_1, lean_object* x_2, 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_WF_GuessLex_mkTupleSyntax___closed__3() { _start: { -lean_object* x_9; -x_9 = l_Lean_Elab_WF_GuessLex_buildTermWF(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_2); -lean_dec(x_1); -return x_9; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("tuple", 5); +return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at_Lean_Elab_WF_GuessLex_trimTermWF___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) { +static lean_object* _init_l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__4() { _start: { -lean_object* x_8; uint8_t x_9; -x_8 = lean_unsigned_to_nat(0u); -x_9 = lean_nat_dec_eq(x_4, x_8); -if (x_9 == 0) -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_10 = lean_unsigned_to_nat(1u); -x_11 = lean_nat_sub(x_4, x_10); -lean_dec(x_4); -x_12 = lean_array_fget(x_3, x_5); -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; -x_14 = lean_ctor_get(x_12, 1); -x_15 = lean_array_get_size(x_14); -x_16 = lean_array_get_size(x_1); -x_17 = lean_nat_dec_lt(x_5, x_16); -lean_dec(x_16); -x_18 = lean_nat_add(x_5, x_10); -if (x_17 == 0) -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; -lean_dec(x_5); -x_19 = l_instInhabitedNat; -x_20 = l___private_Init_Util_0__outOfBounds___rarg(x_19); -x_21 = lean_nat_sub(x_15, x_20); -lean_dec(x_20); -x_22 = l_Array_toSubarray___rarg(x_14, x_21, x_15); -x_23 = l_Array_ofSubarray___rarg(x_22); -x_24 = 0; -lean_ctor_set(x_12, 1, x_23); -lean_ctor_set_uint8(x_12, sizeof(void*)*3, x_24); -x_25 = lean_array_push(x_7, x_12); -x_4 = x_11; -x_5 = x_18; -x_6 = lean_box(0); -x_7 = x_25; -goto _start; -} -else -{ -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_array_fget(x_1, x_5); -lean_dec(x_5); -x_28 = lean_nat_sub(x_15, x_27); -lean_dec(x_27); -x_29 = l_Array_toSubarray___rarg(x_14, x_28, x_15); -x_30 = l_Array_ofSubarray___rarg(x_29); -x_31 = 0; -lean_ctor_set(x_12, 1, x_30); -lean_ctor_set_uint8(x_12, sizeof(void*)*3, x_31); -x_32 = lean_array_push(x_7, x_12); -x_4 = x_11; -x_5 = x_18; -x_6 = lean_box(0); -x_7 = x_32; -goto _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_WF_GuessLex_initFn____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_9____closed__6; +x_2 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__1; +x_3 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__2; +x_4 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__3; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } -else -{ -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; -x_34 = lean_ctor_get(x_12, 0); -x_35 = lean_ctor_get(x_12, 1); -x_36 = lean_ctor_get(x_12, 2); -lean_inc(x_36); -lean_inc(x_35); -lean_inc(x_34); -lean_dec(x_12); -x_37 = lean_array_get_size(x_35); -x_38 = lean_array_get_size(x_1); -x_39 = lean_nat_dec_lt(x_5, x_38); -lean_dec(x_38); -x_40 = lean_nat_add(x_5, x_10); -if (x_39 == 0) -{ -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; lean_object* x_47; lean_object* x_48; -lean_dec(x_5); -x_41 = l_instInhabitedNat; -x_42 = l___private_Init_Util_0__outOfBounds___rarg(x_41); -x_43 = lean_nat_sub(x_37, x_42); -lean_dec(x_42); -x_44 = l_Array_toSubarray___rarg(x_35, x_43, x_37); -x_45 = l_Array_ofSubarray___rarg(x_44); -x_46 = 0; -x_47 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_47, 0, x_34); -lean_ctor_set(x_47, 1, x_45); -lean_ctor_set(x_47, 2, x_36); -lean_ctor_set_uint8(x_47, sizeof(void*)*3, x_46); -x_48 = lean_array_push(x_7, x_47); -x_4 = x_11; -x_5 = x_40; -x_6 = lean_box(0); -x_7 = x_48; -goto _start; } -else +static lean_object* _init_l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__5() { +_start: { -lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; lean_object* x_55; lean_object* x_56; -x_50 = lean_array_fget(x_1, x_5); -lean_dec(x_5); -x_51 = lean_nat_sub(x_37, x_50); -lean_dec(x_50); -x_52 = l_Array_toSubarray___rarg(x_35, x_51, x_37); -x_53 = l_Array_ofSubarray___rarg(x_52); -x_54 = 0; -x_55 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_55, 0, x_34); -lean_ctor_set(x_55, 1, x_53); -lean_ctor_set(x_55, 2, x_36); -lean_ctor_set_uint8(x_55, sizeof(void*)*3, x_54); -x_56 = lean_array_push(x_7, x_55); -x_4 = x_11; -x_5 = x_40; -x_6 = lean_box(0); -x_7 = x_56; -goto _start; -} +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("(", 1); +return x_1; } } -else +static lean_object* _init_l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__6() { +_start: { -lean_dec(x_5); -lean_dec(x_4); -return x_7; -} +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("null", 4); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_trimTermWF(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__7() { _start: { -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_3 = lean_array_get_size(x_2); -x_4 = lean_mk_empty_array_with_capacity(x_3); -x_5 = lean_unsigned_to_nat(0u); -x_6 = l_Array_mapIdxM_map___at_Lean_Elab_WF_GuessLex_trimTermWF___spec__1(x_1, x_2, x_2, x_3, x_5, lean_box(0), x_4); -return x_6; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__6; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at_Lean_Elab_WF_GuessLex_trimTermWF___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_Elab_WF_GuessLex_mkTupleSyntax___closed__8() { _start: { -lean_object* x_8; -x_8 = l_Array_mapIdxM_map___at_Lean_Elab_WF_GuessLex_trimTermWF___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_8; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(",", 1); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_trimTermWF___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_mkTupleSyntax(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_3; -x_3 = l_Lean_Elab_WF_GuessLex_trimTermWF(x_1, x_2); -lean_dec(x_2); -lean_dec(x_1); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: -{ -uint8_t x_15; -x_15 = lean_nat_dec_lt(x_12, x_9); -if (x_15 == 0) -{ -lean_dec(x_12); -lean_dec(x_11); -return x_14; -} -else +lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_7 = lean_array_get_size(x_1); +x_8 = lean_unsigned_to_nat(0u); +x_9 = lean_nat_dec_eq(x_7, x_8); +if (x_9 == 0) { -lean_object* x_16; uint8_t x_17; -x_16 = lean_unsigned_to_nat(0u); -x_17 = lean_nat_dec_eq(x_11, x_16); -if (x_17 == 0) +lean_object* x_10; uint8_t x_11; +x_10 = lean_unsigned_to_nat(1u); +x_11 = lean_nat_dec_eq(x_7, x_10); +if (x_11 == 0) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; lean_object* x_22; lean_object* x_23; -x_18 = lean_unsigned_to_nat(1u); -x_19 = lean_nat_sub(x_11, x_18); -lean_dec(x_11); -x_20 = lean_array_get_size(x_14); -x_21 = lean_nat_dec_lt(x_12, x_20); -lean_dec(x_20); -x_22 = lean_array_fget(x_5, x_12); -x_23 = lean_string_length(x_22); -lean_dec(x_22); -if (x_21 == 0) +lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_12 = lean_ctor_get(x_4, 5); +lean_inc(x_12); +lean_dec(x_4); +x_13 = 0; +x_14 = l_Lean_SourceInfo_fromRef(x_12, x_13); +x_15 = lean_st_ref_get(x_5, x_6); +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) { -lean_object* x_24; lean_object* x_25; uint8_t x_26; -x_24 = l_instInhabitedNat; -x_25 = l___private_Init_Util_0__outOfBounds___rarg(x_24); -x_26 = lean_nat_dec_lt(x_25, x_23); +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_17 = lean_ctor_get(x_15, 0); +lean_dec(x_17); +x_18 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__5; +lean_inc(x_14); +x_19 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_19, 0, x_14); +lean_ctor_set(x_19, 1, x_18); +x_20 = lean_nat_dec_lt(x_8, x_7); +lean_dec(x_7); +x_21 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__8; +lean_inc(x_14); +x_22 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_22, 0, x_14); +lean_ctor_set(x_22, 1, x_21); +x_23 = lean_array_get_size(x_1); +lean_inc(x_1); +x_24 = l_Array_toSubarray___rarg(x_1, x_10, x_23); +x_25 = l_Array_ofSubarray___rarg(x_24); +x_26 = l_Lean_Syntax_SepArray_ofElems(x_21, x_25); lean_dec(x_25); -if (x_26 == 0) +x_27 = l_Lean_Elab_WF_GuessLex_naryVarNames___closed__1; +x_28 = l_Array_append___rarg(x_27, x_26); +x_29 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__7; +lean_inc(x_14); +x_30 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_30, 0, x_14); +lean_ctor_set(x_30, 1, x_29); +lean_ctor_set(x_30, 2, x_28); +x_31 = l_Lean_Elab_WF_GuessLex_collectRecCalls___lambda__2___closed__7; +lean_inc(x_14); +x_32 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_32, 0, x_14); +lean_ctor_set(x_32, 1, x_31); +if (x_20 == 0) { -lean_object* x_27; -lean_dec(x_23); -x_27 = lean_nat_add(x_12, x_10); -lean_dec(x_12); -x_4 = lean_box(0); -x_11 = x_19; -x_12 = x_27; -x_13 = lean_box(0); -goto _start; +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +lean_dec(x_1); +x_33 = lean_box(0); +x_34 = l___private_Init_Util_0__outOfBounds___rarg(x_33); +lean_inc(x_14); +x_35 = l_Lean_Syntax_node3(x_14, x_29, x_34, x_22, x_30); +x_36 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__4; +x_37 = l_Lean_Syntax_node3(x_14, x_36, x_19, x_35, x_32); +lean_ctor_set(x_15, 0, x_37); +return x_15; } else { -lean_object* x_29; lean_object* x_30; -x_29 = lean_array_set(x_14, x_12, x_23); -x_30 = lean_nat_add(x_12, x_10); -lean_dec(x_12); -x_4 = lean_box(0); -x_11 = x_19; -x_12 = x_30; -x_13 = lean_box(0); -x_14 = x_29; -goto _start; +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_38 = lean_array_fget(x_1, x_8); +lean_dec(x_1); +lean_inc(x_14); +x_39 = l_Lean_Syntax_node3(x_14, x_29, x_38, x_22, x_30); +x_40 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__4; +x_41 = l_Lean_Syntax_node3(x_14, x_40, x_19, x_39, x_32); +lean_ctor_set(x_15, 0, x_41); +return x_15; } } else { -lean_object* x_32; uint8_t x_33; -x_32 = lean_array_fget(x_14, x_12); -x_33 = lean_nat_dec_lt(x_32, x_23); -lean_dec(x_32); -if (x_33 == 0) +lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_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; lean_object* x_56; lean_object* x_57; +x_42 = lean_ctor_get(x_15, 1); +lean_inc(x_42); +lean_dec(x_15); +x_43 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__5; +lean_inc(x_14); +x_44 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_44, 0, x_14); +lean_ctor_set(x_44, 1, x_43); +x_45 = lean_nat_dec_lt(x_8, x_7); +lean_dec(x_7); +x_46 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__8; +lean_inc(x_14); +x_47 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_47, 0, x_14); +lean_ctor_set(x_47, 1, x_46); +x_48 = lean_array_get_size(x_1); +lean_inc(x_1); +x_49 = l_Array_toSubarray___rarg(x_1, x_10, x_48); +x_50 = l_Array_ofSubarray___rarg(x_49); +x_51 = l_Lean_Syntax_SepArray_ofElems(x_46, x_50); +lean_dec(x_50); +x_52 = l_Lean_Elab_WF_GuessLex_naryVarNames___closed__1; +x_53 = l_Array_append___rarg(x_52, x_51); +x_54 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__7; +lean_inc(x_14); +x_55 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_55, 0, x_14); +lean_ctor_set(x_55, 1, x_54); +lean_ctor_set(x_55, 2, x_53); +x_56 = l_Lean_Elab_WF_GuessLex_collectRecCalls___lambda__2___closed__7; +lean_inc(x_14); +x_57 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_57, 0, x_14); +lean_ctor_set(x_57, 1, x_56); +if (x_45 == 0) { -lean_object* x_34; -lean_dec(x_23); -x_34 = lean_nat_add(x_12, x_10); -lean_dec(x_12); -x_4 = lean_box(0); -x_11 = x_19; -x_12 = x_34; -x_13 = lean_box(0); -goto _start; +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; +lean_dec(x_1); +x_58 = lean_box(0); +x_59 = l___private_Init_Util_0__outOfBounds___rarg(x_58); +lean_inc(x_14); +x_60 = l_Lean_Syntax_node3(x_14, x_54, x_59, x_47, x_55); +x_61 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__4; +x_62 = l_Lean_Syntax_node3(x_14, x_61, x_44, x_60, x_57); +x_63 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_63, 0, x_62); +lean_ctor_set(x_63, 1, x_42); +return x_63; } else { -lean_object* x_36; lean_object* x_37; -x_36 = lean_array_set(x_14, x_12, x_23); -x_37 = lean_nat_add(x_12, x_10); -lean_dec(x_12); -x_4 = lean_box(0); -x_11 = x_19; -x_12 = x_37; -x_13 = lean_box(0); -x_14 = x_36; -goto _start; +lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_64 = lean_array_fget(x_1, x_8); +lean_dec(x_1); +lean_inc(x_14); +x_65 = l_Lean_Syntax_node3(x_14, x_54, x_64, x_47, x_55); +x_66 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__4; +x_67 = l_Lean_Syntax_node3(x_14, x_66, x_44, x_65, x_57); +x_68 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_68, 0, x_67); +lean_ctor_set(x_68, 1, x_42); +return x_68; } } } else { -lean_dec(x_12); -lean_dec(x_11); -return x_14; -} -} -} +lean_object* x_69; lean_object* x_70; +lean_dec(x_7); +lean_dec(x_4); +x_69 = lean_array_fget(x_1, x_8); +lean_dec(x_1); +x_70 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_70, 0, x_69); +lean_ctor_set(x_70, 1, x_6); +return x_70; } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_formatTable___spec__2(size_t x_1, size_t x_2, lean_object* x_3) { -_start: -{ -uint8_t x_4; -x_4 = lean_usize_dec_lt(x_2, x_1); -if (x_4 == 0) -{ -return x_3; } else { -lean_object* x_5; lean_object* x_6; size_t x_7; size_t x_8; lean_object* x_9; -x_5 = lean_unsigned_to_nat(0u); -x_6 = lean_array_uset(x_3, x_2, x_5); -x_7 = 1; -x_8 = lean_usize_add(x_2, x_7); -x_9 = lean_array_uset(x_6, x_2, x_5); -x_2 = x_8; -x_3 = x_9; -goto _start; -} -} -} -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -uint8_t x_11; -x_11 = lean_nat_dec_lt(x_8, x_5); -if (x_11 == 0) -{ -lean_dec(x_8); +lean_object* x_71; uint8_t x_72; lean_object* x_73; lean_object* x_74; uint8_t x_75; lean_dec(x_7); -return x_10; -} -else -{ -lean_object* x_12; uint8_t x_13; -x_12 = lean_unsigned_to_nat(0u); -x_13 = lean_nat_dec_eq(x_7, x_12); -if (x_13 == 0) +lean_dec(x_1); +x_71 = lean_ctor_get(x_4, 5); +lean_inc(x_71); +lean_dec(x_4); +x_72 = 0; +x_73 = l_Lean_SourceInfo_fromRef(x_71, x_72); +x_74 = lean_st_ref_get(x_5, x_6); +x_75 = !lean_is_exclusive(x_74); +if (x_75 == 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_unsigned_to_nat(1u); -x_15 = lean_nat_sub(x_7, x_14); -lean_dec(x_7); -x_16 = lean_array_fget(x_1, x_8); -x_17 = lean_array_get_size(x_16); -lean_inc(x_17); -x_18 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_18, 0, x_12); -lean_ctor_set(x_18, 1, x_17); -lean_ctor_set(x_18, 2, x_14); -lean_inc(x_17); -x_19 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__1(x_1, x_3, x_8, lean_box(0), x_16, x_17, x_18, x_12, x_17, x_14, x_17, x_12, lean_box(0), x_10); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -x_20 = lean_nat_add(x_8, x_6); -lean_dec(x_8); -x_7 = x_15; -x_8 = x_20; -x_9 = lean_box(0); -x_10 = x_19; -goto _start; +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_76 = lean_ctor_get(x_74, 0); +lean_dec(x_76); +x_77 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__5; +lean_inc(x_73); +x_78 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_78, 0, x_73); +lean_ctor_set(x_78, 1, x_77); +x_79 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__7; +x_80 = l_Lean_Elab_WF_GuessLex_naryVarNames___closed__1; +lean_inc(x_73); +x_81 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_81, 0, x_73); +lean_ctor_set(x_81, 1, x_79); +lean_ctor_set(x_81, 2, x_80); +x_82 = l_Lean_Elab_WF_GuessLex_collectRecCalls___lambda__2___closed__7; +lean_inc(x_73); +x_83 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_83, 0, x_73); +lean_ctor_set(x_83, 1, x_82); +x_84 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__4; +x_85 = l_Lean_Syntax_node3(x_73, x_84, x_78, x_81, x_83); +lean_ctor_set(x_74, 0, x_85); +return x_74; } else { -lean_dec(x_8); -lean_dec(x_7); -return x_10; +lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; +x_86 = lean_ctor_get(x_74, 1); +lean_inc(x_86); +lean_dec(x_74); +x_87 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__5; +lean_inc(x_73); +x_88 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_88, 0, x_73); +lean_ctor_set(x_88, 1, x_87); +x_89 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__7; +x_90 = l_Lean_Elab_WF_GuessLex_naryVarNames___closed__1; +lean_inc(x_73); +x_91 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_91, 0, x_73); +lean_ctor_set(x_91, 1, x_89); +lean_ctor_set(x_91, 2, x_90); +x_92 = l_Lean_Elab_WF_GuessLex_collectRecCalls___lambda__2___closed__7; +lean_inc(x_73); +x_93 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_93, 0, x_73); +lean_ctor_set(x_93, 1, x_92); +x_94 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__4; +x_95 = l_Lean_Syntax_node3(x_73, x_94, x_88, x_91, x_93); +x_96 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_96, 0, x_95); +lean_ctor_set(x_96, 1, x_86); +return x_96; } } } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__4___closed__1() { +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_mkTupleSyntax___boxed(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(" ", 1); -return x_1; +lean_object* x_7; +x_7 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +return x_7; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_formatTable___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_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__1(size_t x_1, size_t x_2, lean_object* x_3) { _start: { -uint8_t x_6; -x_6 = lean_nat_dec_le(x_3, x_2); -if (x_6 == 0) -{ -lean_object* x_7; uint8_t x_8; -x_7 = lean_unsigned_to_nat(0u); -x_8 = lean_nat_dec_eq(x_1, 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_13; -x_9 = lean_unsigned_to_nat(1u); -x_10 = lean_nat_sub(x_1, x_9); -lean_dec(x_1); -x_11 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__4___closed__1; -x_12 = lean_string_append(x_5, x_11); -x_13 = lean_nat_add(x_2, x_4); -lean_dec(x_2); -x_1 = x_10; -x_2 = x_13; -x_5 = x_12; -goto _start; -} -else +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) { -lean_dec(x_2); -lean_dec(x_1); -return x_5; -} +return x_3; } else { -lean_dec(x_2); -lean_dec(x_1); -return x_5; +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = lean_mk_syntax_ident(x_5); +x_9 = 1; +x_10 = lean_usize_add(x_2, x_9); +x_11 = lean_array_uset(x_7, x_2, x_8); +x_2 = x_10; +x_3 = x_11; +goto _start; } } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__5___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__4___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_5 = lean_unsigned_to_nat(1u); -x_6 = lean_nat_add(x_1, x_5); -x_7 = lean_nat_dec_lt(x_6, x_2); -lean_dec(x_6); -if (x_7 == 0) +lean_object* x_11; lean_object* x_12; +lean_inc(x_3); +x_11 = l_Lean_resolveGlobalName___at_Lean_Elab_WF_GuessLex_naryVarNames_freshen___spec__1(x_3, x_6, x_7, x_8, x_9, x_10); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +if (lean_obj_tag(x_12) == 0) { -lean_object* x_8; -x_8 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_8, 0, x_3); -return x_8; -} -else +uint8_t x_13; +x_13 = !lean_is_exclusive(x_11); +if (x_13 == 0) { -lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_9 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__4___closed__1; -x_10 = lean_string_append(x_3, x_9); -x_11 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_11, 0, x_10); +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_14 = lean_ctor_get(x_11, 0); +lean_dec(x_14); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_3); +lean_ctor_set(x_15, 1, x_4); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_1); +lean_ctor_set(x_16, 1, x_15); +x_17 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_11, 0, x_17); return x_11; } +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_11, 1); +lean_inc(x_18); +lean_dec(x_11); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_3); +lean_ctor_set(x_19, 1, x_4); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_1); +lean_ctor_set(x_20, 1, x_19); +x_21 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_21, 0, 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_18); +return x_22; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___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: +else { -lean_object* x_7; lean_object* x_8; uint8_t x_9; -x_7 = lean_string_append(x_5, x_1); -x_8 = lean_unsigned_to_nat(0u); -x_9 = lean_nat_dec_eq(x_2, x_8); -if (x_9 == 0) +lean_object* x_23; lean_object* x_24; +x_23 = lean_ctor_get(x_12, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_12, 1); +lean_inc(x_24); +lean_dec(x_12); +if (lean_obj_tag(x_24) == 0) { -lean_object* x_10; lean_object* x_11; -lean_dec(x_4); -lean_dec(x_1); -x_10 = lean_box(0); -x_11 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__5___lambda__1(x_2, x_3, x_7, x_10); -lean_dec(x_3); -lean_dec(x_2); +uint8_t x_25; +x_25 = !lean_is_exclusive(x_11); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_26 = lean_ctor_get(x_11, 0); +lean_dec(x_26); +x_27 = lean_ctor_get(x_23, 0); +lean_inc(x_27); +lean_dec(x_23); +x_28 = lean_name_eq(x_27, x_2); +lean_dec(x_27); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +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_1); +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_11, 0, x_31); return x_11; } else { -lean_object* x_12; uint8_t x_13; lean_object* x_14; -x_12 = lean_array_get_size(x_4); -x_13 = lean_nat_dec_lt(x_2, x_12); -lean_dec(x_12); -x_14 = lean_string_length(x_1); +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_dec(x_1); -if (x_13 == 0) -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -lean_dec(x_4); -x_15 = l_instInhabitedNat; -x_16 = l___private_Init_Util_0__outOfBounds___rarg(x_15); -x_17 = lean_nat_sub(x_16, x_14); -lean_dec(x_14); -lean_dec(x_16); -x_18 = lean_unsigned_to_nat(1u); -lean_inc(x_17); -x_19 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__4(x_17, x_8, x_17, x_18, x_7); -lean_dec(x_17); -x_20 = lean_box(0); -x_21 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__5___lambda__1(x_2, x_3, x_19, x_20); -lean_dec(x_3); -lean_dec(x_2); -return x_21; +lean_inc(x_3); +x_32 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_32, 0, x_3); +x_33 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_33, 0, x_32); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_3); +lean_ctor_set(x_34, 1, x_4); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +x_36 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_11, 0, x_36); +return x_11; +} } 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_array_fget(x_4, x_2); -lean_dec(x_4); -x_23 = lean_nat_sub(x_22, x_14); -lean_dec(x_14); -lean_dec(x_22); -x_24 = lean_unsigned_to_nat(1u); -lean_inc(x_23); -x_25 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__4(x_23, x_8, x_23, x_24, x_7); +lean_object* x_37; lean_object* x_38; uint8_t x_39; +x_37 = lean_ctor_get(x_11, 1); +lean_inc(x_37); +lean_dec(x_11); +x_38 = lean_ctor_get(x_23, 0); +lean_inc(x_38); lean_dec(x_23); -x_26 = lean_box(0); -x_27 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__5___lambda__1(x_2, x_3, x_25, x_26); -lean_dec(x_3); -lean_dec(x_2); -return x_27; -} -} -} -} -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { -_start: -{ -uint8_t x_16; -x_16 = lean_nat_dec_lt(x_13, x_10); -if (x_16 == 0) +x_39 = lean_name_eq(x_38, x_2); +lean_dec(x_38); +if (x_39 == 0) { -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_3); -return x_15; +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_3); +lean_ctor_set(x_40, 1, x_4); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_1); +lean_ctor_set(x_41, 1, x_40); +x_42 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_42, 0, x_41); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_37); +return x_43; } else { -lean_object* x_17; uint8_t x_18; -x_17 = lean_unsigned_to_nat(0u); -x_18 = lean_nat_dec_eq(x_12, x_17); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_19 = lean_unsigned_to_nat(1u); -x_20 = lean_nat_sub(x_12, x_19); -lean_dec(x_12); -x_21 = lean_array_fget(x_6, x_13); -x_22 = lean_nat_dec_lt(x_17, x_13); -if (x_22 == 0) -{ -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_23 = lean_box(0); +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +lean_dec(x_1); lean_inc(x_3); -lean_inc(x_7); -lean_inc(x_13); -x_24 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__5___lambda__2(x_21, x_13, x_7, x_3, x_15, x_23); -x_25 = lean_ctor_get(x_24, 0); -lean_inc(x_25); -lean_dec(x_24); -x_26 = lean_nat_add(x_13, x_11); -lean_dec(x_13); -x_5 = lean_box(0); -x_12 = x_20; -x_13 = x_26; -x_14 = lean_box(0); -x_15 = x_25; -goto _start; +x_44 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_44, 0, x_3); +x_45 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_45, 0, x_44); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_3); +lean_ctor_set(x_46, 1, x_4); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +x_48 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_48, 0, 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_37); +return x_49; +} +} } else { -lean_object* x_28; uint8_t x_29; lean_object* x_30; -x_28 = lean_array_get_size(x_3); -x_29 = lean_nat_dec_lt(x_13, x_28); -lean_dec(x_28); -x_30 = lean_string_length(x_21); -if (x_29 == 0) +uint8_t x_50; +lean_dec(x_24); +lean_dec(x_23); +x_50 = !lean_is_exclusive(x_11); +if (x_50 == 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; -x_31 = l_instInhabitedNat; -x_32 = l___private_Init_Util_0__outOfBounds___rarg(x_31); -x_33 = lean_nat_sub(x_32, x_30); -lean_dec(x_30); -lean_dec(x_32); -lean_inc(x_33); -x_34 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__4(x_33, x_17, x_33, x_19, x_15); -lean_dec(x_33); -x_35 = lean_box(0); -lean_inc(x_3); -lean_inc(x_7); -lean_inc(x_13); -x_36 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__5___lambda__2(x_21, x_13, x_7, x_3, x_34, x_35); -x_37 = lean_ctor_get(x_36, 0); -lean_inc(x_37); -lean_dec(x_36); -x_38 = lean_nat_add(x_13, x_11); -lean_dec(x_13); -x_5 = lean_box(0); -x_12 = x_20; -x_13 = x_38; -x_14 = lean_box(0); -x_15 = x_37; -goto _start; +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_51 = lean_ctor_get(x_11, 0); +lean_dec(x_51); +x_52 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_52, 0, x_3); +lean_ctor_set(x_52, 1, x_4); +x_53 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_53, 0, x_1); +lean_ctor_set(x_53, 1, x_52); +x_54 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_54, 0, x_53); +lean_ctor_set(x_11, 0, x_54); +return x_11; } 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; -x_40 = lean_array_fget(x_3, x_13); -x_41 = lean_nat_sub(x_40, x_30); -lean_dec(x_30); -lean_dec(x_40); -lean_inc(x_41); -x_42 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__4(x_41, x_17, x_41, x_19, x_15); -lean_dec(x_41); -x_43 = lean_box(0); -lean_inc(x_3); -lean_inc(x_7); -lean_inc(x_13); -x_44 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__5___lambda__2(x_21, x_13, x_7, x_3, x_42, x_43); -x_45 = lean_ctor_get(x_44, 0); -lean_inc(x_45); -lean_dec(x_44); -x_46 = lean_nat_add(x_13, x_11); -lean_dec(x_13); -x_5 = lean_box(0); -x_12 = x_20; -x_13 = x_46; -x_14 = lean_box(0); -x_15 = x_45; -goto _start; -} -} +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_55 = lean_ctor_get(x_11, 1); +lean_inc(x_55); +lean_dec(x_11); +x_56 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_56, 0, x_3); +lean_ctor_set(x_56, 1, x_4); +x_57 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_57, 0, x_1); +lean_ctor_set(x_57, 1, x_56); +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_55); +return x_59; } -else -{ -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_3); -return x_15; } } } } -static lean_object* _init_l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__6___closed__1() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__4___closed__1() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("\n", 1); -return x_1; +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_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -uint8_t x_12; -x_12 = lean_nat_dec_lt(x_9, x_6); -if (x_12 == 0) +uint8_t x_13; +x_13 = lean_nat_dec_le(x_5, x_4); +if (x_13 == 0) { -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_4); -return x_11; -} -else -{ -lean_object* x_13; uint8_t x_14; -x_13 = lean_unsigned_to_nat(0u); -x_14 = lean_nat_dec_eq(x_8, x_13); -if (x_14 == 0) +lean_object* x_14; uint8_t x_15; +x_14 = lean_unsigned_to_nat(0u); +x_15 = lean_nat_dec_eq(x_3, x_14); +if (x_15 == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_15 = lean_unsigned_to_nat(1u); -x_16 = lean_nat_sub(x_8, x_15); -lean_dec(x_8); -x_17 = lean_array_fget(x_1, x_9); -x_18 = lean_array_get_size(x_17); +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_16 = lean_unsigned_to_nat(1u); +x_17 = lean_nat_sub(x_3, x_16); +lean_dec(x_3); +x_18 = lean_ctor_get(x_7, 1); lean_inc(x_18); -x_19 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_19, 0, x_13); -lean_ctor_set(x_19, 1, x_18); -lean_ctor_set(x_19, 2, x_15); -lean_inc_n(x_18, 2); -lean_inc(x_4); -x_20 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__5(x_1, x_3, x_4, x_9, lean_box(0), x_17, x_18, x_19, x_13, x_18, x_15, x_18, x_13, lean_box(0), x_11); -lean_dec(x_19); -lean_dec(x_18); +lean_dec(x_7); +x_19 = lean_ctor_get(x_18, 1); +lean_inc(x_19); +if (lean_obj_tag(x_19) == 0) +{ +uint8_t x_20; lean_dec(x_17); -x_21 = lean_nat_add(x_9, x_15); -x_22 = lean_nat_dec_lt(x_21, x_2); -lean_dec(x_21); -if (x_22 == 0) +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +x_20 = !lean_is_exclusive(x_18); +if (x_20 == 0) { -lean_object* x_23; -x_23 = lean_nat_add(x_9, x_7); -lean_dec(x_9); -x_8 = x_16; -x_9 = x_23; -x_10 = lean_box(0); -x_11 = x_20; -goto _start; +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_21 = lean_ctor_get(x_18, 1); +lean_dec(x_21); +x_22 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__4___closed__1; +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_18); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_12); +return x_24; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__6___closed__1; -x_26 = lean_string_append(x_20, x_25); -x_27 = lean_nat_add(x_9, x_7); -lean_dec(x_9); -x_8 = x_16; -x_9 = x_27; -x_10 = lean_box(0); -x_11 = x_26; -goto _start; +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_18, 0); +lean_inc(x_25); +lean_dec(x_18); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_19); +x_27 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__4___closed__1; +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_26); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_12); +return x_29; } } else { -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_4); -return x_11; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_formatTable(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; uint8_t x_4; size_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_2 = lean_array_get_size(x_1); -x_3 = lean_unsigned_to_nat(0u); -x_4 = lean_nat_dec_lt(x_3, x_2); -x_5 = 0; -x_6 = lean_unsigned_to_nat(1u); +lean_object* x_30; 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_18, 0); +lean_inc(x_30); +lean_dec(x_18); +x_31 = lean_ctor_get(x_19, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_19, 1); +lean_inc(x_32); +lean_dec(x_19); +x_33 = l_Lean_Name_append(x_31, x_30); +x_34 = lean_box(0); +lean_inc(x_10); lean_inc(x_2); -x_7 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_7, 0, x_3); -lean_ctor_set(x_7, 1, x_2); -lean_ctor_set(x_7, 2, x_6); -if (x_4 == 0) +x_35 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__4___lambda__1(x_2, x_1, x_33, x_32, x_34, x_8, x_9, x_10, x_11, x_12); +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +if (lean_obj_tag(x_36) == 0) { -lean_object* x_16; lean_object* x_17; -x_16 = l_Lean_Elab_WF_GuessLex_naryVarNames___closed__1; -x_17 = l___private_Init_Util_0__outOfBounds___rarg(x_16); -x_8 = x_17; -goto block_15; +uint8_t x_37; +lean_dec(x_17); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +x_37 = !lean_is_exclusive(x_35); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_35, 0); +lean_dec(x_38); +x_39 = lean_ctor_get(x_36, 0); +lean_inc(x_39); +lean_dec(x_36); +lean_ctor_set(x_35, 0, x_39); +return x_35; } else { -lean_object* x_18; -x_18 = lean_array_fget(x_1, x_3); -x_8 = x_18; -goto block_15; +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_35, 1); +lean_inc(x_40); +lean_dec(x_35); +x_41 = lean_ctor_get(x_36, 0); +lean_inc(x_41); +lean_dec(x_36); +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; } -block_15: +} +else { -lean_object* x_9; size_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_9 = lean_array_get_size(x_8); -x_10 = lean_usize_of_nat(x_9); -lean_dec(x_9); -x_11 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_formatTable___spec__2(x_10, x_5, x_8); -lean_inc(x_2); -x_12 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__3(x_1, x_2, x_7, x_3, x_2, x_6, x_2, x_3, lean_box(0), x_11); -x_13 = l_Lean_Elab_WF_GuessLex_initFn____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_9____closed__3; -lean_inc(x_2); -x_14 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__6(x_1, x_2, x_7, x_12, x_3, x_2, x_6, x_2, x_3, lean_box(0), x_13); -lean_dec(x_7); -lean_dec(x_2); -return x_14; +lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_43 = lean_ctor_get(x_35, 1); +lean_inc(x_43); +lean_dec(x_35); +x_44 = lean_ctor_get(x_36, 0); +lean_inc(x_44); +lean_dec(x_36); +x_45 = lean_nat_add(x_4, x_6); +lean_dec(x_4); +x_3 = x_17; +x_4 = x_45; +x_7 = x_44; +x_12 = x_43; +goto _start; } } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___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: +else { -lean_object* x_15; -x_15 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_object* x_47; 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_15; -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_formatTable___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -size_t x_4; size_t x_5; lean_object* x_6; -x_4 = lean_unbox_usize(x_1); -lean_dec(x_1); -x_5 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_formatTable___spec__2(x_4, x_5, x_3); -return x_6; +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_7); +lean_ctor_set(x_47, 1, x_12); +return x_47; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___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 { -lean_object* x_11; -x_11 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_48; +lean_dec(x_10); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -return x_11; +x_48 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_48, 0, x_7); +lean_ctor_set(x_48, 1, x_12); +return x_48; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +} +LEAN_EXPORT lean_object* l_Lean_unresolveNameGlobal_unresolveNameCore___at_Lean_Elab_WF_GuessLex_buildTermWF___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) { _start: { -lean_object* x_6; -x_6 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__4(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_6; +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_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__5___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_unresolveNameGlobal_unresolveNameCore___at_Lean_Elab_WF_GuessLex_buildTermWF___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) { _start: { -lean_object* x_5; -x_5 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__5___lambda__1(x_1, x_2, x_3, x_4); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___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_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_8 = l_Lean_Name_componentsRev(x_2); +x_9 = lean_unsigned_to_nat(0u); +x_10 = l_List_lengthTRAux___rarg(x_8, x_9); +x_11 = lean_box(0); +x_12 = lean_box(0); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_8); +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_unsigned_to_nat(1u); +lean_inc(x_10); +x_16 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__4(x_1, x_11, x_10, x_9, x_10, x_15, x_14, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_10); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +lean_dec(x_17); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_7; -x_7 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__5___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_6); -return x_7; -} -} -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { -_start: +uint8_t x_19; +x_19 = !lean_is_exclusive(x_16); +if (x_19 == 0) { -lean_object* x_16; -x_16 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); +lean_object* x_20; +x_20 = lean_ctor_get(x_16, 0); +lean_dec(x_20); +lean_ctor_set(x_16, 0, x_11); return x_16; } -} -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: +else { -lean_object* x_12; -x_12 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_12; +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_16, 1); +lean_inc(x_21); +lean_dec(x_16); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_11); +lean_ctor_set(x_22, 1, x_21); +return x_22; } } -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_formatTable___boxed(lean_object* x_1) { -_start: +else { -lean_object* x_2; -x_2 = l_Lean_Elab_WF_GuessLex_formatTable(x_1); -lean_dec(x_1); -return x_2; -} +uint8_t x_23; +x_23 = !lean_is_exclusive(x_16); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_16, 0); +lean_dec(x_24); +x_25 = lean_ctor_get(x_18, 0); +lean_inc(x_25); +lean_dec(x_18); +lean_ctor_set(x_16, 0, x_25); +return x_16; } -static lean_object* _init_l_Lean_Elab_WF_GuessLex_RecCallWithContext_posString___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes(":", 1); -return x_1; +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_16, 1); +lean_inc(x_26); +lean_dec(x_16); +x_27 = lean_ctor_get(x_18, 0); +lean_inc(x_27); +lean_dec(x_18); +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; } } -static lean_object* _init_l_Lean_Elab_WF_GuessLex_RecCallWithContext_posString___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("-", 1); -return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_RecCallWithContext_posString(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_buildTermWF___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_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_7; lean_object* x_8; uint8_t x_9; lean_object* x_10; -x_7 = lean_ctor_get(x_4, 1); -x_8 = lean_ctor_get(x_1, 0); -lean_inc(x_8); -lean_dec(x_1); -x_9 = 0; -x_10 = l_Lean_Syntax_getPos_x3f(x_8, x_9); -if (lean_obj_tag(x_10) == 0) +uint8_t x_12; +x_12 = lean_usize_dec_lt(x_5, x_4); +if (x_12 == 0) { -lean_object* x_11; lean_object* x_12; -lean_dec(x_8); -x_11 = l_Lean_Elab_WF_GuessLex_initFn____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_9____closed__3; -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_6); -return x_12; +lean_object* x_13; +lean_dec(x_9); +lean_dec(x_2); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_6); +lean_ctor_set(x_13, 1, x_11); +return x_13; } else { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_13 = lean_ctor_get(x_10, 0); -lean_inc(x_13); -lean_dec(x_10); -x_14 = l_Lean_FileMap_toPosition(x_7, x_13); -lean_dec(x_13); -x_15 = l_Lean_Syntax_getTailPos_x3f(x_8, x_9); -x_16 = lean_ctor_get(x_14, 0); -lean_inc(x_16); +lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_dec(x_6); +x_14 = lean_array_uget(x_3, x_5); +lean_inc(x_9); +x_15 = l_Lean_unresolveNameGlobal_unresolveNameCore___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__3(x_1, x_14, x_7, x_8, x_9, x_10, x_11); +x_16 = lean_ctor_get(x_15, 0); lean_inc(x_16); -x_17 = l_Nat_repr(x_16); -x_18 = l_Lean_Elab_WF_GuessLex_initFn____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_9____closed__3; -x_19 = lean_string_append(x_18, x_17); -lean_dec(x_17); -x_20 = l_Lean_Elab_WF_GuessLex_RecCallWithContext_posString___closed__1; -x_21 = lean_string_append(x_19, x_20); -x_22 = lean_ctor_get(x_14, 1); -lean_inc(x_22); -lean_dec(x_14); -x_23 = l_Nat_repr(x_22); -x_24 = lean_string_append(x_21, x_23); -lean_dec(x_23); -x_25 = lean_string_append(x_24, x_18); -if (lean_obj_tag(x_15) == 0) +if (lean_obj_tag(x_16) == 0) { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -lean_dec(x_16); -x_26 = lean_string_append(x_25, x_18); -x_27 = lean_string_append(x_26, x_18); -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; +lean_object* x_17; size_t x_18; size_t x_19; +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_5, x_18); +lean_inc(x_2); +{ +size_t _tmp_4 = x_19; +lean_object* _tmp_5 = x_2; +lean_object* _tmp_10 = x_17; +x_5 = _tmp_4; +x_6 = _tmp_5; +x_11 = _tmp_10; +} +goto _start; } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; -x_29 = lean_ctor_get(x_15, 0); -lean_inc(x_29); -lean_dec(x_15); -x_30 = l_Lean_FileMap_toPosition(x_7, x_29); -lean_dec(x_29); -x_31 = lean_ctor_get(x_30, 0); -lean_inc(x_31); -x_32 = lean_nat_dec_eq(x_31, x_16); -lean_dec(x_16); -if (x_32 == 0) +uint8_t x_21; +lean_dec(x_9); +lean_dec(x_2); +x_21 = !lean_is_exclusive(x_15); +if (x_21 == 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; -x_33 = l_Nat_repr(x_31); -x_34 = l_Lean_Elab_WF_GuessLex_RecCallWithContext_posString___closed__2; -x_35 = lean_string_append(x_34, x_33); -lean_dec(x_33); -x_36 = lean_string_append(x_35, x_20); -x_37 = lean_ctor_get(x_30, 1); -lean_inc(x_37); -lean_dec(x_30); -x_38 = l_Nat_repr(x_37); -x_39 = lean_string_append(x_36, x_38); -lean_dec(x_38); -x_40 = lean_string_append(x_39, x_18); -x_41 = lean_string_append(x_25, x_40); -lean_dec(x_40); -x_42 = lean_string_append(x_41, 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_6); -return x_43; +lean_object* x_22; uint8_t x_23; +x_22 = lean_ctor_get(x_15, 0); +lean_dec(x_22); +x_23 = !lean_is_exclusive(x_16); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; +x_24 = lean_box(0); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_16); +lean_ctor_set(x_25, 1, x_24); +lean_ctor_set(x_15, 0, x_25); +return x_15; } 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_dec(x_31); -x_44 = lean_ctor_get(x_30, 1); -lean_inc(x_44); -lean_dec(x_30); -x_45 = l_Nat_repr(x_44); -x_46 = l_Lean_Elab_WF_GuessLex_RecCallWithContext_posString___closed__2; -x_47 = lean_string_append(x_46, x_45); -lean_dec(x_45); -x_48 = lean_string_append(x_47, x_18); -x_49 = lean_string_append(x_25, x_48); -lean_dec(x_48); -x_50 = lean_string_append(x_49, x_18); -x_51 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_51, 0, x_50); -lean_ctor_set(x_51, 1, x_6); -return x_51; -} -} -} +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_26 = lean_ctor_get(x_16, 0); +lean_inc(x_26); +lean_dec(x_16); +x_27 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_27, 0, x_26); +x_28 = lean_box(0); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +lean_ctor_set(x_15, 0, x_29); +return x_15; } } -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_RecCallWithContext_posString___boxed(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_Elab_WF_GuessLex_RecCallWithContext_posString(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_30; 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_15, 1); +lean_inc(x_30); +lean_dec(x_15); +x_31 = lean_ctor_get(x_16, 0); +lean_inc(x_31); +if (lean_is_exclusive(x_16)) { + lean_ctor_release(x_16, 0); + x_32 = x_16; +} else { + lean_dec_ref(x_16); + x_32 = lean_box(0); } +if (lean_is_scalar(x_32)) { + x_33 = lean_alloc_ctor(1, 1, 0); +} else { + x_33 = x_32; } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_explainNonMutualFailure___spec__1(size_t x_1, size_t x_2, lean_object* x_3) { -_start: -{ -uint8_t x_4; -x_4 = lean_usize_dec_lt(x_2, x_1); -if (x_4 == 0) -{ -return x_3; +lean_ctor_set(x_33, 0, x_31); +x_34 = lean_box(0); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_30); +return x_36; } -else -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; lean_object* x_10; size_t x_11; size_t x_12; lean_object* x_13; -x_5 = lean_array_uget(x_3, x_2); -x_6 = lean_unsigned_to_nat(0u); -x_7 = lean_array_uset(x_3, x_2, x_6); -x_8 = lean_erase_macro_scopes(x_5); -x_9 = 1; -x_10 = l_Lean_Name_toString(x_8, x_9); -x_11 = 1; -x_12 = lean_usize_add(x_2, x_11); -x_13 = lean_array_uset(x_7, x_2, x_10); -x_2 = x_12; -x_3 = x_13; -goto _start; } } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainNonMutualFailure___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +} +LEAN_EXPORT lean_object* l_Lean_unresolveNameGlobal___at_Lean_Elab_WF_GuessLex_buildTermWF___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) { _start: { -uint8_t x_12; -x_12 = lean_nat_dec_le(x_4, x_3); -if (x_12 == 0) +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; size_t x_20; size_t x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +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, 0); +lean_inc(x_11); +lean_dec(x_9); +lean_inc(x_1); +x_12 = l_Lean_getRevAliases(x_11, x_1); +x_13 = l_List_redLength___rarg(x_12); +x_14 = lean_mk_empty_array_with_capacity(x_13); +lean_dec(x_13); +x_15 = l_List_toArrayAux___rarg(x_12, x_14); +x_16 = l_Lean_rootNamespace; +lean_inc(x_1); +x_17 = l_Lean_Name_append(x_16, x_1); +x_18 = lean_array_push(x_15, 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_Lean_Elab_WF_GuessLex_evalRecCall___lambda__2___closed__4; +x_23 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__5(x_1, x_22, x_18, x_20, x_21, x_22, x_3, x_4, x_5, x_6, x_10); +lean_dec(x_18); +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +lean_dec(x_24); +if (lean_obj_tag(x_25) == 0) { -lean_object* x_13; uint8_t x_14; -x_13 = lean_unsigned_to_nat(0u); -x_14 = lean_nat_dec_eq(x_2, x_13); -if (x_14 == 0) +uint8_t x_26; +x_26 = !lean_is_exclusive(x_23); +if (x_26 == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_15 = lean_unsigned_to_nat(1u); -x_16 = lean_nat_sub(x_2, x_15); -lean_dec(x_2); -x_17 = l_Lean_Elab_WF_GuessLex_RecCallCache_prettyEntry(x_1, x_3, x_3, x_7, x_8, x_9, x_10, x_11); -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_17, 1); -lean_inc(x_19); -lean_dec(x_17); -x_20 = lean_array_push(x_6, x_18); -x_21 = lean_nat_add(x_3, x_5); -lean_dec(x_3); -x_2 = x_16; -x_3 = x_21; -x_6 = x_20; -x_11 = x_19; -goto _start; +lean_object* x_27; +x_27 = lean_ctor_get(x_23, 0); +lean_dec(x_27); +lean_ctor_set(x_23, 0, x_1); +return x_23; } else { -lean_object* x_23; -lean_dec(x_3); -lean_dec(x_2); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_6); -lean_ctor_set(x_23, 1, x_11); -return x_23; +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_23, 1); +lean_inc(x_28); +lean_dec(x_23); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_1); +lean_ctor_set(x_29, 1, x_28); +return x_29; } } else { -lean_object* x_24; -lean_dec(x_3); -lean_dec(x_2); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_6); -lean_ctor_set(x_24, 1, x_11); -return x_24; +uint8_t x_30; +lean_dec(x_1); +x_30 = !lean_is_exclusive(x_23); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; +x_31 = lean_ctor_get(x_23, 0); +lean_dec(x_31); +x_32 = lean_ctor_get(x_25, 0); +lean_inc(x_32); +lean_dec(x_25); +lean_ctor_set(x_23, 0, x_32); +return x_23; } +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_23, 1); +lean_inc(x_33); +lean_dec(x_23); +x_34 = lean_ctor_get(x_25, 0); +lean_inc(x_34); +lean_dec(x_25); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_33); +return x_35; } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainNonMutualFailure___spec__3___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes(") ", 2); -return x_1; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainNonMutualFailure___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_unresolveNameGlobal___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__2___lambda__2(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -uint8_t x_13; -x_13 = lean_nat_dec_le(x_5, x_4); -if (x_13 == 0) -{ -lean_object* x_14; uint8_t x_15; -x_14 = lean_unsigned_to_nat(0u); -x_15 = lean_nat_dec_eq(x_3, x_14); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_16 = lean_unsigned_to_nat(1u); -x_17 = lean_nat_sub(x_3, x_16); lean_dec(x_3); -x_18 = !lean_is_exclusive(x_7); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_19 = lean_ctor_get(x_7, 0); -x_20 = lean_ctor_get(x_7, 1); -x_21 = lean_ctor_get(x_19, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_19, 1); -lean_inc(x_22); -x_23 = lean_ctor_get(x_19, 2); -lean_inc(x_23); -x_24 = lean_nat_dec_lt(x_22, x_23); -if (x_24 == 0) +if (x_1 == 0) { -lean_object* x_25; -lean_dec(x_23); -lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_17); +lean_object* x_9; lean_object* x_10; +x_9 = lean_box(0); +x_10 = l_Lean_unresolveNameGlobal___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__2___lambda__1(x_2, x_9, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_5); lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_7); -lean_ctor_set(x_25, 1, x_12); -return x_25; +return x_10; } else { -uint8_t x_26; -x_26 = !lean_is_exclusive(x_19); -if (x_26 == 0) +lean_object* x_11; lean_object* x_12; +lean_inc(x_2); +x_11 = l_Lean_resolveGlobalName___at_Lean_Elab_WF_GuessLex_naryVarNames_freshen___spec__1(x_2, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +if (lean_obj_tag(x_12) == 0) +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_11); +if (x_13 == 0) +{ +lean_object* x_14; +x_14 = lean_ctor_get(x_11, 0); +lean_dec(x_14); +lean_ctor_set(x_11, 0, x_2); +return x_11; +} +else +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_11, 1); +lean_inc(x_15); +lean_dec(x_11); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_2); +lean_ctor_set(x_16, 1, x_15); +return x_16; +} +} +else +{ +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_12, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_12, 1); +lean_inc(x_18); +lean_dec(x_12); +if (lean_obj_tag(x_18) == 0) +{ +uint8_t x_19; +x_19 = !lean_is_exclusive(x_11); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_11, 0); +lean_dec(x_20); +x_21 = lean_ctor_get(x_17, 0); +lean_inc(x_21); +lean_dec(x_17); +lean_inc(x_21); +x_22 = lean_private_to_user_name(x_21); +if (lean_obj_tag(x_22) == 0) +{ +uint8_t x_23; +x_23 = lean_name_eq(x_21, x_2); +lean_dec(x_21); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; +x_24 = l_Lean_rootNamespace; +x_25 = l_Lean_Name_append(x_24, x_2); +lean_ctor_set(x_11, 0, x_25); +return x_11; +} +else +{ +lean_ctor_set(x_11, 0, x_2); +return x_11; +} +} +else +{ +lean_object* x_26; uint8_t x_27; +lean_dec(x_21); +x_26 = lean_ctor_get(x_22, 0); +lean_inc(x_26); +lean_dec(x_22); +x_27 = lean_name_eq(x_26, x_2); +lean_dec(x_26); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; +x_28 = l_Lean_rootNamespace; +x_29 = l_Lean_Name_append(x_28, x_2); +lean_ctor_set(x_11, 0, x_29); +return x_11; +} +else +{ +lean_ctor_set(x_11, 0, x_2); +return x_11; +} +} +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_11, 1); +lean_inc(x_30); +lean_dec(x_11); +x_31 = lean_ctor_get(x_17, 0); +lean_inc(x_31); +lean_dec(x_17); +lean_inc(x_31); +x_32 = lean_private_to_user_name(x_31); +if (lean_obj_tag(x_32) == 0) +{ +uint8_t x_33; +x_33 = lean_name_eq(x_31, x_2); +lean_dec(x_31); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = l_Lean_rootNamespace; +x_35 = l_Lean_Name_append(x_34, x_2); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_30); +return x_36; +} +else +{ +lean_object* x_37; +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_2); +lean_ctor_set(x_37, 1, x_30); +return x_37; +} +} +else +{ +lean_object* x_38; uint8_t x_39; +lean_dec(x_31); +x_38 = lean_ctor_get(x_32, 0); +lean_inc(x_38); +lean_dec(x_32); +x_39 = lean_name_eq(x_38, x_2); +lean_dec(x_38); +if (x_39 == 0) +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = l_Lean_rootNamespace; +x_41 = l_Lean_Name_append(x_40, x_2); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_30); +return x_42; +} +else +{ +lean_object* x_43; +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_2); +lean_ctor_set(x_43, 1, x_30); +return x_43; +} +} +} +} +else +{ +uint8_t x_44; +lean_dec(x_18); +lean_dec(x_17); +x_44 = !lean_is_exclusive(x_11); +if (x_44 == 0) +{ +lean_object* x_45; +x_45 = lean_ctor_get(x_11, 0); +lean_dec(x_45); +lean_ctor_set(x_11, 0, x_2); +return x_11; +} +else +{ +lean_object* x_46; lean_object* x_47; +x_46 = lean_ctor_get(x_11, 1); +lean_inc(x_46); +lean_dec(x_11); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_2); +lean_ctor_set(x_47, 1, x_46); +return x_47; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_unresolveNameGlobal___at_Lean_Elab_WF_GuessLex_buildTermWF___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) { +_start: +{ +uint8_t x_8; +x_8 = l_Lean_Name_hasMacroScopes(x_1); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; +x_9 = lean_box(0); +x_10 = l_Lean_unresolveNameGlobal___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__2___lambda__2(x_2, x_1, x_9, x_3, x_4, x_5, x_6, x_7); +return x_10; +} +else +{ +lean_object* x_11; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_1); +lean_ctor_set(x_11, 1, x_7); +return x_11; +} +} +} +static lean_object* _init_l_Array_anyMUnsafe_any___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__6___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_WF_GuessLex_mkSizeOf___closed__3; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__6(lean_object* x_1, size_t x_2, size_t x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_usize_dec_eq(x_2, x_3); +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_Elab_WF_GuessLex_buildTermWF___spec__6___closed__1; +x_7 = lean_name_eq(x_5, x_6); +lean_dec(x_5); +if (x_7 == 0) +{ +size_t x_8; size_t x_9; +x_8 = 1; +x_9 = lean_usize_add(x_2, x_8); +x_2 = x_9; +goto _start; +} +else +{ +uint8_t x_11; +x_11 = 1; +return x_11; +} +} +else +{ +uint8_t x_12; +x_12 = 0; +return x_12; +} +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("app", 3); +return x_1; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7___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_WF_GuessLex_initFn____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_9____closed__6; +x_2 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__1; +x_3 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__2; +x_4 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7___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_Lean_Elab_WF_GuessLex_buildTermWF___spec__7___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_WF_GuessLex_mkSizeOf___closed__4; +x_2 = lean_mk_syntax_ident(x_1); +return x_2; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("num", 3); +return x_1; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7___closed__4; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("0", 1); +return x_1; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("1", 1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, lean_object* x_6, size_t x_7, size_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +uint8_t x_15; +x_15 = lean_usize_dec_lt(x_8, x_7); +if (x_15 == 0) +{ +lean_object* x_16; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_9); +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_array_uget(x_9, x_8); +x_18 = lean_unsigned_to_nat(0u); +x_19 = lean_array_uset(x_9, x_8, x_18); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; lean_object* x_26; +x_20 = lean_ctor_get(x_17, 0); +lean_inc(x_20); +lean_dec(x_17); +x_21 = lean_array_get_size(x_20); +x_22 = lean_nat_dec_lt(x_4, x_21); +lean_dec(x_21); +x_23 = lean_array_get_size(x_6); +x_24 = lean_array_get_size(x_3); +x_25 = lean_nat_dec_lt(x_4, x_24); +lean_dec(x_24); +if (x_22 == 0) +{ +lean_object* x_98; lean_object* x_99; +lean_dec(x_20); +x_98 = l_instInhabitedNat; +x_99 = l___private_Init_Util_0__outOfBounds___rarg(x_98); +x_26 = x_99; +goto block_97; +} +else +{ +lean_object* x_100; +x_100 = lean_array_fget(x_20, x_4); +lean_dec(x_20); +x_26 = x_100; +goto block_97; +} +block_97: +{ +lean_object* x_27; uint8_t x_93; +x_93 = lean_nat_dec_lt(x_26, x_23); +lean_dec(x_23); +if (x_93 == 0) +{ +lean_object* x_94; lean_object* x_95; +x_94 = lean_box(0); +x_95 = l___private_Init_Util_0__outOfBounds___rarg(x_94); +x_27 = x_95; +goto block_92; +} +else +{ +lean_object* x_96; +x_96 = lean_array_fget(x_6, x_26); +x_27 = x_96; +goto block_92; +} +block_92: +{ +lean_object* x_28; +if (x_25 == 0) +{ +lean_object* x_73; lean_object* x_74; uint8_t x_75; +x_73 = l_Lean_Elab_WF_GuessLex_naryVarNames___closed__1; +x_74 = l___private_Init_Util_0__outOfBounds___rarg(x_73); +x_75 = l_Array_contains___at___private_Lean_Meta_FunInfo_0__Lean_Meta_collectDeps_visit___spec__2(x_74, x_26); +lean_dec(x_26); +lean_dec(x_74); +if (x_75 == 0) +{ +lean_object* x_76; +x_76 = lean_box(0); +x_28 = x_76; +goto block_72; +} +else +{ +lean_object* x_77; lean_object* x_78; size_t x_79; size_t x_80; lean_object* x_81; +x_77 = lean_st_ref_get(x_13, x_14); +x_78 = lean_ctor_get(x_77, 1); +lean_inc(x_78); +lean_dec(x_77); +x_79 = 1; +x_80 = lean_usize_add(x_8, x_79); +x_81 = lean_array_uset(x_19, x_8, x_27); +x_8 = x_80; +x_9 = x_81; +x_14 = x_78; +goto _start; +} +} +else +{ +lean_object* x_83; uint8_t x_84; +x_83 = lean_array_fget(x_3, x_4); +x_84 = l_Array_contains___at___private_Lean_Meta_FunInfo_0__Lean_Meta_collectDeps_visit___spec__2(x_83, x_26); +lean_dec(x_26); +lean_dec(x_83); +if (x_84 == 0) +{ +lean_object* x_85; +x_85 = lean_box(0); +x_28 = x_85; +goto block_72; +} +else +{ +lean_object* x_86; lean_object* x_87; size_t x_88; size_t x_89; lean_object* x_90; +x_86 = lean_st_ref_get(x_13, x_14); +x_87 = lean_ctor_get(x_86, 1); +lean_inc(x_87); +lean_dec(x_86); +x_88 = 1; +x_89 = lean_usize_add(x_8, x_88); +x_90 = lean_array_uset(x_19, x_8, x_27); +x_8 = x_89; +x_9 = x_90; +x_14 = x_87; +goto _start; +} +} +block_72: +{ +lean_object* x_29; uint8_t x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_51; +lean_dec(x_28); +x_29 = l_Lean_Elab_WF_GuessLex_mkSizeOf___closed__4; +x_30 = 0; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_31 = l_Lean_unresolveNameGlobal___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__2(x_29, x_30, x_10, x_11, x_12, x_13, x_14); +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 = lean_array_get_size(x_1); +x_35 = lean_nat_dec_lt(x_4, x_34); +lean_dec(x_34); +x_36 = lean_ctor_get(x_12, 5); +lean_inc(x_36); +x_37 = l_Lean_SourceInfo_fromRef(x_36, x_30); +x_38 = lean_st_ref_get(x_13, x_33); +if (x_35 == 0) +{ +lean_object* x_69; lean_object* x_70; +x_69 = l_Lean_Elab_WF_GuessLex_naryVarNames___closed__1; +x_70 = l___private_Init_Util_0__outOfBounds___rarg(x_69); +x_51 = x_70; +goto block_68; +} +else +{ +lean_object* x_71; +x_71 = lean_array_fget(x_1, x_4); +x_51 = x_71; +goto block_68; +} +block_50: +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; size_t x_46; size_t x_47; lean_object* x_48; +lean_dec(x_39); +x_40 = lean_ctor_get(x_38, 1); +lean_inc(x_40); +lean_dec(x_38); +x_41 = lean_mk_syntax_ident(x_32); +x_42 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__7; +lean_inc(x_37); +x_43 = l_Lean_Syntax_node1(x_37, x_42, x_27); +x_44 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7___closed__2; +x_45 = l_Lean_Syntax_node2(x_37, x_44, x_41, x_43); +x_46 = 1; +x_47 = lean_usize_add(x_8, x_46); +x_48 = lean_array_uset(x_19, x_8, x_45); +x_8 = x_47; +x_9 = x_48; +x_14 = x_40; +goto _start; +} +block_68: +{ +lean_object* x_52; uint8_t x_53; +x_52 = lean_array_get_size(x_51); +x_53 = lean_nat_dec_lt(x_18, x_52); +if (x_53 == 0) +{ +lean_object* x_54; +lean_dec(x_52); +lean_dec(x_51); +x_54 = lean_box(0); +x_39 = x_54; +goto block_50; +} +else +{ +size_t x_55; uint8_t x_56; +x_55 = lean_usize_of_nat(x_52); +lean_dec(x_52); +x_56 = l_Array_anyMUnsafe_any___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__6(x_51, x_5, x_55); +lean_dec(x_51); +if (x_56 == 0) +{ +lean_object* x_57; +x_57 = lean_box(0); +x_39 = x_57; +goto block_50; +} +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; size_t x_64; size_t x_65; lean_object* x_66; +lean_dec(x_32); +x_58 = lean_ctor_get(x_38, 1); +lean_inc(x_58); +lean_dec(x_38); +x_59 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__7; +lean_inc(x_37); +x_60 = l_Lean_Syntax_node1(x_37, x_59, x_27); +x_61 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7___closed__2; +x_62 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7___closed__3; +x_63 = l_Lean_Syntax_node2(x_37, x_61, x_62, x_60); +x_64 = 1; +x_65 = lean_usize_add(x_8, x_64); +x_66 = lean_array_uset(x_19, x_8, x_63); +x_8 = x_65; +x_9 = x_66; +x_14 = x_58; +goto _start; +} +} +} +} +} +} +} +else +{ +lean_object* x_101; uint8_t x_102; +x_101 = lean_ctor_get(x_17, 0); +lean_inc(x_101); +lean_dec(x_17); +x_102 = lean_nat_dec_eq(x_101, x_4); +lean_dec(x_101); +if (x_102 == 0) +{ +lean_object* x_103; uint8_t x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; size_t x_112; size_t x_113; lean_object* x_114; +x_103 = lean_ctor_get(x_12, 5); +lean_inc(x_103); +x_104 = 0; +x_105 = l_Lean_SourceInfo_fromRef(x_103, x_104); +x_106 = lean_st_ref_get(x_13, x_14); +x_107 = lean_ctor_get(x_106, 1); +lean_inc(x_107); +lean_dec(x_106); +x_108 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7___closed__6; +lean_inc(x_105); +x_109 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_109, 0, x_105); +lean_ctor_set(x_109, 1, x_108); +x_110 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7___closed__5; +x_111 = l_Lean_Syntax_node1(x_105, x_110, x_109); +x_112 = 1; +x_113 = lean_usize_add(x_8, x_112); +x_114 = lean_array_uset(x_19, x_8, x_111); +x_8 = x_113; +x_9 = x_114; +x_14 = x_107; +goto _start; +} +else +{ +lean_object* x_116; uint8_t x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; size_t x_125; size_t x_126; lean_object* x_127; +x_116 = lean_ctor_get(x_12, 5); +lean_inc(x_116); +x_117 = 0; +x_118 = l_Lean_SourceInfo_fromRef(x_116, x_117); +x_119 = lean_st_ref_get(x_13, x_14); +x_120 = lean_ctor_get(x_119, 1); +lean_inc(x_120); +lean_dec(x_119); +x_121 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7___closed__7; +lean_inc(x_118); +x_122 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_122, 0, x_118); +lean_ctor_set(x_122, 1, x_121); +x_123 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7___closed__5; +x_124 = l_Lean_Syntax_node1(x_118, x_123, x_122); +x_125 = 1; +x_126 = lean_usize_add(x_8, x_125); +x_127 = lean_array_uset(x_19, x_8, x_124); +x_8 = x_126; +x_9 = x_127; +x_14 = x_120; +goto _start; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__8(size_t x_1, size_t x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) +{ +return x_3; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = 1; +x_9 = lean_usize_add(x_2, x_8); +x_10 = lean_array_uset(x_7, x_2, x_5); +x_2 = x_9; +x_3 = x_10; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; uint8_t x_16; +x_15 = lean_unsigned_to_nat(0u); +x_16 = lean_nat_dec_eq(x_6, x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; size_t x_21; size_t x_22; lean_object* x_23; lean_object* x_24; size_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; size_t 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; +x_17 = lean_unsigned_to_nat(1u); +x_18 = lean_nat_sub(x_6, x_17); +lean_dec(x_6); +x_19 = lean_array_fget(x_5, x_7); +x_20 = lean_array_get_size(x_19); +x_21 = lean_usize_of_nat(x_20); +lean_dec(x_20); +x_22 = 0; +x_23 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__1(x_21, x_22, x_19); +x_24 = lean_array_get_size(x_4); +x_25 = lean_usize_of_nat(x_24); +lean_dec(x_24); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_4); +x_26 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7(x_1, x_2, x_3, x_7, x_22, x_23, x_25, x_22, x_4, x_10, x_11, x_12, x_13, x_14); +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_12); +x_29 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax(x_27, x_10, x_11, x_12, x_13, 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_array_get_size(x_23); +x_33 = lean_usize_of_nat(x_32); +lean_dec(x_32); +x_34 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__8(x_33, x_22, x_23); +x_35 = lean_box(0); +x_36 = 1; +x_37 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_34); +lean_ctor_set(x_37, 2, x_30); +lean_ctor_set_uint8(x_37, sizeof(void*)*3, x_36); +x_38 = lean_nat_add(x_7, x_17); +lean_dec(x_7); +x_39 = lean_array_push(x_9, x_37); +x_6 = x_18; +x_7 = x_38; +x_8 = lean_box(0); +x_9 = x_39; +x_14 = x_31; +goto _start; +} +else +{ +lean_object* x_41; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_9); +lean_ctor_set(x_41, 1, x_14); +return x_41; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_buildTermWF(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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; +x_10 = lean_array_get_size(x_2); +x_11 = lean_mk_empty_array_with_capacity(x_10); +x_12 = lean_unsigned_to_nat(0u); +x_13 = l_Array_mapIdxM_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__9(x_1, x_2, x_3, x_4, x_2, x_10, x_12, lean_box(0), x_11, x_5, x_6, x_7, x_8, x_9); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__1(x_4, x_5, x_3); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__4___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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_Lean_Elab_WF_GuessLex_buildTermWF___spec__4___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_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_buildTermWF___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_13; +x_13 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_1); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_unresolveNameGlobal_unresolveNameCore___at_Lean_Elab_WF_GuessLex_buildTermWF___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) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_unresolveNameGlobal_unresolveNameCore___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__3___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_Lean_unresolveNameGlobal_unresolveNameCore___at_Lean_Elab_WF_GuessLex_buildTermWF___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) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_unresolveNameGlobal_unresolveNameCore___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_buildTermWF___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) { +_start: +{ +size_t x_12; size_t x_13; lean_object* x_14; +x_12 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_13 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_14 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__5(x_1, x_2, x_3, x_12, x_13, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_3); +lean_dec(x_1); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_unresolveNameGlobal___at_Lean_Elab_WF_GuessLex_buildTermWF___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) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_unresolveNameGlobal___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__2___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_unresolveNameGlobal___at_Lean_Elab_WF_GuessLex_buildTermWF___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) { +_start: +{ +uint8_t x_9; lean_object* x_10; +x_9 = lean_unbox(x_1); +lean_dec(x_1); +x_10 = l_Lean_unresolveNameGlobal___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__2___lambda__2(x_9, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_unresolveNameGlobal___at_Lean_Elab_WF_GuessLex_buildTermWF___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) { +_start: +{ +uint8_t x_8; lean_object* x_9; +x_8 = lean_unbox(x_2); +lean_dec(x_2); +x_9 = l_Lean_unresolveNameGlobal___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__2(x_1, x_8, x_3, x_4, x_5, x_6, x_7); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__6___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; +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_Elab_WF_GuessLex_buildTermWF___spec__6(x_1, x_4, x_5); +lean_dec(x_1); +x_7 = lean_box(x_6); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___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: +{ +size_t x_15; size_t x_16; size_t x_17; lean_object* x_18; +x_15 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_16 = lean_unbox_usize(x_7); +lean_dec(x_7); +x_17 = lean_unbox_usize(x_8); +lean_dec(x_8); +x_18 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__7(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); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__8(x_4, x_5, x_3); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; +x_15 = l_Array_mapIdxM_map___at_Lean_Elab_WF_GuessLex_buildTermWF___spec__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_buildTermWF___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_WF_GuessLex_buildTermWF(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at_Lean_Elab_WF_GuessLex_trimTermWF___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_8; uint8_t x_9; +x_8 = lean_unsigned_to_nat(0u); +x_9 = lean_nat_dec_eq(x_4, x_8); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_10 = lean_unsigned_to_nat(1u); +x_11 = lean_nat_sub(x_4, x_10); +lean_dec(x_4); +x_12 = lean_array_fget(x_3, x_5); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; +x_14 = lean_ctor_get(x_12, 1); +x_15 = lean_array_get_size(x_14); +x_16 = lean_array_get_size(x_1); +x_17 = lean_nat_dec_lt(x_5, x_16); +lean_dec(x_16); +x_18 = lean_nat_add(x_5, x_10); +if (x_17 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; +lean_dec(x_5); +x_19 = l_instInhabitedNat; +x_20 = l___private_Init_Util_0__outOfBounds___rarg(x_19); +x_21 = lean_nat_sub(x_15, x_20); +lean_dec(x_20); +x_22 = l_Array_toSubarray___rarg(x_14, x_21, x_15); +x_23 = l_Array_ofSubarray___rarg(x_22); +x_24 = 0; +lean_ctor_set(x_12, 1, x_23); +lean_ctor_set_uint8(x_12, sizeof(void*)*3, x_24); +x_25 = lean_array_push(x_7, x_12); +x_4 = x_11; +x_5 = x_18; +x_6 = lean_box(0); +x_7 = x_25; +goto _start; +} +else +{ +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_array_fget(x_1, x_5); +lean_dec(x_5); +x_28 = lean_nat_sub(x_15, x_27); +lean_dec(x_27); +x_29 = l_Array_toSubarray___rarg(x_14, x_28, x_15); +x_30 = l_Array_ofSubarray___rarg(x_29); +x_31 = 0; +lean_ctor_set(x_12, 1, x_30); +lean_ctor_set_uint8(x_12, sizeof(void*)*3, x_31); +x_32 = lean_array_push(x_7, x_12); +x_4 = x_11; +x_5 = x_18; +x_6 = lean_box(0); +x_7 = x_32; +goto _start; +} +} +else +{ +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; +x_34 = lean_ctor_get(x_12, 0); +x_35 = lean_ctor_get(x_12, 1); +x_36 = lean_ctor_get(x_12, 2); +lean_inc(x_36); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_12); +x_37 = lean_array_get_size(x_35); +x_38 = lean_array_get_size(x_1); +x_39 = lean_nat_dec_lt(x_5, x_38); +lean_dec(x_38); +x_40 = lean_nat_add(x_5, x_10); +if (x_39 == 0) +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; lean_object* x_47; lean_object* x_48; +lean_dec(x_5); +x_41 = l_instInhabitedNat; +x_42 = l___private_Init_Util_0__outOfBounds___rarg(x_41); +x_43 = lean_nat_sub(x_37, x_42); +lean_dec(x_42); +x_44 = l_Array_toSubarray___rarg(x_35, x_43, x_37); +x_45 = l_Array_ofSubarray___rarg(x_44); +x_46 = 0; +x_47 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_47, 0, x_34); +lean_ctor_set(x_47, 1, x_45); +lean_ctor_set(x_47, 2, x_36); +lean_ctor_set_uint8(x_47, sizeof(void*)*3, x_46); +x_48 = lean_array_push(x_7, x_47); +x_4 = x_11; +x_5 = x_40; +x_6 = lean_box(0); +x_7 = x_48; +goto _start; +} +else +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; lean_object* x_55; lean_object* x_56; +x_50 = lean_array_fget(x_1, x_5); +lean_dec(x_5); +x_51 = lean_nat_sub(x_37, x_50); +lean_dec(x_50); +x_52 = l_Array_toSubarray___rarg(x_35, x_51, x_37); +x_53 = l_Array_ofSubarray___rarg(x_52); +x_54 = 0; +x_55 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_55, 0, x_34); +lean_ctor_set(x_55, 1, x_53); +lean_ctor_set(x_55, 2, x_36); +lean_ctor_set_uint8(x_55, sizeof(void*)*3, x_54); +x_56 = lean_array_push(x_7, x_55); +x_4 = x_11; +x_5 = x_40; +x_6 = lean_box(0); +x_7 = x_56; +goto _start; +} +} +} +else +{ +lean_dec(x_5); +lean_dec(x_4); +return x_7; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_trimTermWF(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_2); +x_4 = lean_mk_empty_array_with_capacity(x_3); +x_5 = lean_unsigned_to_nat(0u); +x_6 = l_Array_mapIdxM_map___at_Lean_Elab_WF_GuessLex_trimTermWF___spec__1(x_1, x_2, x_2, x_3, x_5, lean_box(0), x_4); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at_Lean_Elab_WF_GuessLex_trimTermWF___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_8; +x_8 = l_Array_mapIdxM_map___at_Lean_Elab_WF_GuessLex_trimTermWF___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_trimTermWF___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Elab_WF_GuessLex_trimTermWF(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +uint8_t x_15; +x_15 = lean_nat_dec_lt(x_12, x_9); +if (x_15 == 0) +{ +lean_dec(x_12); +lean_dec(x_11); +return x_14; +} +else +{ +lean_object* x_16; uint8_t x_17; +x_16 = lean_unsigned_to_nat(0u); +x_17 = lean_nat_dec_eq(x_11, x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; lean_object* x_22; lean_object* x_23; +x_18 = lean_unsigned_to_nat(1u); +x_19 = lean_nat_sub(x_11, x_18); +lean_dec(x_11); +x_20 = lean_array_get_size(x_14); +x_21 = lean_nat_dec_lt(x_12, x_20); +lean_dec(x_20); +x_22 = lean_array_fget(x_5, x_12); +x_23 = lean_string_length(x_22); +lean_dec(x_22); +if (x_21 == 0) +{ +lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_24 = l_instInhabitedNat; +x_25 = l___private_Init_Util_0__outOfBounds___rarg(x_24); +x_26 = lean_nat_dec_lt(x_25, x_23); +lean_dec(x_25); +if (x_26 == 0) +{ +lean_object* x_27; +lean_dec(x_23); +x_27 = lean_nat_add(x_12, x_10); +lean_dec(x_12); +x_4 = lean_box(0); +x_11 = x_19; +x_12 = x_27; +x_13 = lean_box(0); +goto _start; +} +else +{ +lean_object* x_29; lean_object* x_30; +x_29 = lean_array_set(x_14, x_12, x_23); +x_30 = lean_nat_add(x_12, x_10); +lean_dec(x_12); +x_4 = lean_box(0); +x_11 = x_19; +x_12 = x_30; +x_13 = lean_box(0); +x_14 = x_29; +goto _start; +} +} +else +{ +lean_object* x_32; uint8_t x_33; +x_32 = lean_array_fget(x_14, x_12); +x_33 = lean_nat_dec_lt(x_32, x_23); +lean_dec(x_32); +if (x_33 == 0) +{ +lean_object* x_34; +lean_dec(x_23); +x_34 = lean_nat_add(x_12, x_10); +lean_dec(x_12); +x_4 = lean_box(0); +x_11 = x_19; +x_12 = x_34; +x_13 = lean_box(0); +goto _start; +} +else +{ +lean_object* x_36; lean_object* x_37; +x_36 = lean_array_set(x_14, x_12, x_23); +x_37 = lean_nat_add(x_12, x_10); +lean_dec(x_12); +x_4 = lean_box(0); +x_11 = x_19; +x_12 = x_37; +x_13 = lean_box(0); +x_14 = x_36; +goto _start; +} +} +} +else +{ +lean_dec(x_12); +lean_dec(x_11); +return x_14; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_formatTable___spec__2(size_t x_1, size_t x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) +{ +return x_3; +} +else +{ +lean_object* x_5; lean_object* x_6; size_t x_7; size_t x_8; lean_object* x_9; +x_5 = lean_unsigned_to_nat(0u); +x_6 = lean_array_uset(x_3, x_2, x_5); +x_7 = 1; +x_8 = lean_usize_add(x_2, x_7); +x_9 = lean_array_uset(x_6, x_2, x_5); +x_2 = x_8; +x_3 = x_9; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +uint8_t x_11; +x_11 = lean_nat_dec_lt(x_8, x_5); +if (x_11 == 0) +{ +lean_dec(x_8); +lean_dec(x_7); +return x_10; +} +else +{ +lean_object* x_12; uint8_t x_13; +x_12 = lean_unsigned_to_nat(0u); +x_13 = lean_nat_dec_eq(x_7, 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; +x_14 = lean_unsigned_to_nat(1u); +x_15 = lean_nat_sub(x_7, x_14); +lean_dec(x_7); +x_16 = lean_array_fget(x_1, x_8); +x_17 = lean_array_get_size(x_16); +lean_inc(x_17); +x_18 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_18, 0, x_12); +lean_ctor_set(x_18, 1, x_17); +lean_ctor_set(x_18, 2, x_14); +lean_inc(x_17); +x_19 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__1(x_1, x_3, x_8, lean_box(0), x_16, x_17, x_18, x_12, x_17, x_14, x_17, x_12, lean_box(0), x_10); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +x_20 = lean_nat_add(x_8, x_6); +lean_dec(x_8); +x_7 = x_15; +x_8 = x_20; +x_9 = lean_box(0); +x_10 = x_19; +goto _start; +} +else +{ +lean_dec(x_8); +lean_dec(x_7); +return x_10; +} +} +} +} +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__4___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(" ", 1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__4(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; +x_6 = lean_nat_dec_le(x_3, x_2); +if (x_6 == 0) +{ +lean_object* x_7; uint8_t x_8; +x_7 = lean_unsigned_to_nat(0u); +x_8 = lean_nat_dec_eq(x_1, 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_13; +x_9 = lean_unsigned_to_nat(1u); +x_10 = lean_nat_sub(x_1, x_9); +lean_dec(x_1); +x_11 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__4___closed__1; +x_12 = lean_string_append(x_5, x_11); +x_13 = lean_nat_add(x_2, x_4); +lean_dec(x_2); +x_1 = x_10; +x_2 = x_13; +x_5 = x_12; +goto _start; +} +else +{ +lean_dec(x_2); +lean_dec(x_1); +return x_5; +} +} +else +{ +lean_dec(x_2); +lean_dec(x_1); +return x_5; +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___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; uint8_t x_7; +x_5 = lean_unsigned_to_nat(1u); +x_6 = lean_nat_add(x_1, x_5); +x_7 = lean_nat_dec_lt(x_6, x_2); +lean_dec(x_6); +if (x_7 == 0) +{ +lean_object* x_8; +x_8 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_8, 0, x_3); +return x_8; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_9 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__4___closed__1; +x_10 = lean_string_append(x_3, x_9); +x_11 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_11, 0, x_10); +return x_11; +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___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; lean_object* x_8; uint8_t x_9; +x_7 = lean_string_append(x_5, x_1); +x_8 = lean_unsigned_to_nat(0u); +x_9 = lean_nat_dec_eq(x_2, x_8); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; +lean_dec(x_4); +lean_dec(x_1); +x_10 = lean_box(0); +x_11 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__5___lambda__1(x_2, x_3, x_7, x_10); +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} +else +{ +lean_object* x_12; uint8_t x_13; lean_object* x_14; +x_12 = lean_array_get_size(x_4); +x_13 = lean_nat_dec_lt(x_2, x_12); +lean_dec(x_12); +x_14 = lean_string_length(x_1); +lean_dec(x_1); +if (x_13 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +lean_dec(x_4); +x_15 = l_instInhabitedNat; +x_16 = l___private_Init_Util_0__outOfBounds___rarg(x_15); +x_17 = lean_nat_sub(x_16, x_14); +lean_dec(x_14); +lean_dec(x_16); +x_18 = lean_unsigned_to_nat(1u); +lean_inc(x_17); +x_19 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__4(x_17, x_8, x_17, x_18, x_7); +lean_dec(x_17); +x_20 = lean_box(0); +x_21 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__5___lambda__1(x_2, x_3, x_19, x_20); +lean_dec(x_3); +lean_dec(x_2); +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_array_fget(x_4, x_2); +lean_dec(x_4); +x_23 = lean_nat_sub(x_22, x_14); +lean_dec(x_14); +lean_dec(x_22); +x_24 = lean_unsigned_to_nat(1u); +lean_inc(x_23); +x_25 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__4(x_23, x_8, x_23, x_24, x_7); +lean_dec(x_23); +x_26 = lean_box(0); +x_27 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__5___lambda__1(x_2, x_3, x_25, x_26); +lean_dec(x_3); +lean_dec(x_2); +return x_27; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +uint8_t x_16; +x_16 = lean_nat_dec_lt(x_13, x_10); +if (x_16 == 0) +{ +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_7); +lean_dec(x_3); +return x_15; +} +else +{ +lean_object* x_17; uint8_t x_18; +x_17 = lean_unsigned_to_nat(0u); +x_18 = lean_nat_dec_eq(x_12, x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_19 = lean_unsigned_to_nat(1u); +x_20 = lean_nat_sub(x_12, x_19); +lean_dec(x_12); +x_21 = lean_array_fget(x_6, x_13); +x_22 = lean_nat_dec_lt(x_17, x_13); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_23 = lean_box(0); +lean_inc(x_3); +lean_inc(x_7); +lean_inc(x_13); +x_24 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__5___lambda__2(x_21, x_13, x_7, x_3, x_15, x_23); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +lean_dec(x_24); +x_26 = lean_nat_add(x_13, x_11); +lean_dec(x_13); +x_5 = lean_box(0); +x_12 = x_20; +x_13 = x_26; +x_14 = lean_box(0); +x_15 = x_25; +goto _start; +} +else +{ +lean_object* x_28; uint8_t x_29; lean_object* x_30; +x_28 = lean_array_get_size(x_3); +x_29 = lean_nat_dec_lt(x_13, x_28); +lean_dec(x_28); +x_30 = lean_string_length(x_21); +if (x_29 == 0) +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_31 = l_instInhabitedNat; +x_32 = l___private_Init_Util_0__outOfBounds___rarg(x_31); +x_33 = lean_nat_sub(x_32, x_30); +lean_dec(x_30); +lean_dec(x_32); +lean_inc(x_33); +x_34 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__4(x_33, x_17, x_33, x_19, x_15); +lean_dec(x_33); +x_35 = lean_box(0); +lean_inc(x_3); +lean_inc(x_7); +lean_inc(x_13); +x_36 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__5___lambda__2(x_21, x_13, x_7, x_3, x_34, x_35); +x_37 = lean_ctor_get(x_36, 0); +lean_inc(x_37); +lean_dec(x_36); +x_38 = lean_nat_add(x_13, x_11); +lean_dec(x_13); +x_5 = lean_box(0); +x_12 = x_20; +x_13 = x_38; +x_14 = lean_box(0); +x_15 = x_37; +goto _start; +} +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; +x_40 = lean_array_fget(x_3, x_13); +x_41 = lean_nat_sub(x_40, x_30); +lean_dec(x_30); +lean_dec(x_40); +lean_inc(x_41); +x_42 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__4(x_41, x_17, x_41, x_19, x_15); +lean_dec(x_41); +x_43 = lean_box(0); +lean_inc(x_3); +lean_inc(x_7); +lean_inc(x_13); +x_44 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__5___lambda__2(x_21, x_13, x_7, x_3, x_42, x_43); +x_45 = lean_ctor_get(x_44, 0); +lean_inc(x_45); +lean_dec(x_44); +x_46 = lean_nat_add(x_13, x_11); +lean_dec(x_13); +x_5 = lean_box(0); +x_12 = x_20; +x_13 = x_46; +x_14 = lean_box(0); +x_15 = x_45; +goto _start; +} +} +} +else +{ +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_7); +lean_dec(x_3); +return x_15; +} +} +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__6___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("\n", 1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +uint8_t x_12; +x_12 = lean_nat_dec_lt(x_9, x_6); +if (x_12 == 0) +{ +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_4); +return x_11; +} +else +{ +lean_object* x_13; uint8_t x_14; +x_13 = lean_unsigned_to_nat(0u); +x_14 = lean_nat_dec_eq(x_8, x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_15 = lean_unsigned_to_nat(1u); +x_16 = lean_nat_sub(x_8, x_15); +lean_dec(x_8); +x_17 = lean_array_fget(x_1, x_9); +x_18 = lean_array_get_size(x_17); +lean_inc(x_18); +x_19 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_19, 0, x_13); +lean_ctor_set(x_19, 1, x_18); +lean_ctor_set(x_19, 2, x_15); +lean_inc_n(x_18, 2); +lean_inc(x_4); +x_20 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__5(x_1, x_3, x_4, x_9, lean_box(0), x_17, x_18, x_19, x_13, x_18, x_15, x_18, x_13, lean_box(0), x_11); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +x_21 = lean_nat_add(x_9, x_15); +x_22 = lean_nat_dec_lt(x_21, x_2); +lean_dec(x_21); +if (x_22 == 0) +{ +lean_object* x_23; +x_23 = lean_nat_add(x_9, x_7); +lean_dec(x_9); +x_8 = x_16; +x_9 = x_23; +x_10 = lean_box(0); +x_11 = x_20; +goto _start; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__6___closed__1; +x_26 = lean_string_append(x_20, x_25); +x_27 = lean_nat_add(x_9, x_7); +lean_dec(x_9); +x_8 = x_16; +x_9 = x_27; +x_10 = lean_box(0); +x_11 = x_26; +goto _start; +} +} +else +{ +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_4); +return x_11; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_formatTable(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; uint8_t x_4; size_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_2 = lean_array_get_size(x_1); +x_3 = lean_unsigned_to_nat(0u); +x_4 = lean_nat_dec_lt(x_3, x_2); +x_5 = 0; +x_6 = lean_unsigned_to_nat(1u); +lean_inc(x_2); +x_7 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_7, 0, x_3); +lean_ctor_set(x_7, 1, x_2); +lean_ctor_set(x_7, 2, x_6); +if (x_4 == 0) +{ +lean_object* x_16; lean_object* x_17; +x_16 = l_Lean_Elab_WF_GuessLex_naryVarNames___closed__1; +x_17 = l___private_Init_Util_0__outOfBounds___rarg(x_16); +x_8 = x_17; +goto block_15; +} +else +{ +lean_object* x_18; +x_18 = lean_array_fget(x_1, x_3); +x_8 = x_18; +goto block_15; +} +block_15: +{ +lean_object* x_9; size_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_9 = lean_array_get_size(x_8); +x_10 = lean_usize_of_nat(x_9); +lean_dec(x_9); +x_11 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_formatTable___spec__2(x_10, x_5, x_8); +lean_inc(x_2); +x_12 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__3(x_1, x_2, x_7, x_3, x_2, x_6, x_2, x_3, lean_box(0), x_11); +x_13 = l_Lean_Elab_WF_GuessLex_initFn____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_9____closed__3; +lean_inc(x_2); +x_14 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__6(x_1, x_2, x_7, x_12, x_3, x_2, x_6, x_2, x_3, lean_box(0), x_13); +lean_dec(x_7); +lean_dec(x_2); +return x_14; +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; +x_15 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_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); +lean_dec(x_1); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_formatTable___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_formatTable___spec__2(x_4, x_5, x_3); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___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: +{ +lean_object* x_11; +x_11 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +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_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__4___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_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__4(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_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___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_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__5___lambda__1(x_1, x_2, x_3, x_4); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___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_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__5___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_6); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +lean_object* x_16; +x_16 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_formatTable___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Elab_WF_GuessLex_formatTable(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Elab_WF_GuessLex_RecCallWithContext_posString___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(":", 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_WF_GuessLex_RecCallWithContext_posString___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("-", 1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_RecCallWithContext_posString(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_10; +x_7 = lean_ctor_get(x_4, 1); +x_8 = lean_ctor_get(x_1, 0); +lean_inc(x_8); +lean_dec(x_1); +x_9 = 0; +x_10 = l_Lean_Syntax_getPos_x3f(x_8, x_9); +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_11; lean_object* x_12; +lean_dec(x_8); +x_11 = l_Lean_Elab_WF_GuessLex_initFn____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_9____closed__3; +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_6); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_13 = lean_ctor_get(x_10, 0); +lean_inc(x_13); +lean_dec(x_10); +x_14 = l_Lean_FileMap_toPosition(x_7, x_13); +lean_dec(x_13); +x_15 = l_Lean_Syntax_getTailPos_x3f(x_8, x_9); +x_16 = lean_ctor_get(x_14, 0); +lean_inc(x_16); +lean_inc(x_16); +x_17 = l_Nat_repr(x_16); +x_18 = l_Lean_Elab_WF_GuessLex_initFn____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_9____closed__3; +x_19 = lean_string_append(x_18, x_17); +lean_dec(x_17); +x_20 = l_Lean_Elab_WF_GuessLex_RecCallWithContext_posString___closed__1; +x_21 = lean_string_append(x_19, x_20); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_dec(x_14); +x_23 = l_Nat_repr(x_22); +x_24 = lean_string_append(x_21, x_23); +lean_dec(x_23); +x_25 = lean_string_append(x_24, x_18); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +lean_dec(x_16); +x_26 = lean_string_append(x_25, x_18); +x_27 = lean_string_append(x_26, x_18); +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; +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_29 = lean_ctor_get(x_15, 0); +lean_inc(x_29); +lean_dec(x_15); +x_30 = l_Lean_FileMap_toPosition(x_7, x_29); +lean_dec(x_29); +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_nat_dec_eq(x_31, x_16); +lean_dec(x_16); +if (x_32 == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_33 = l_Nat_repr(x_31); +x_34 = l_Lean_Elab_WF_GuessLex_RecCallWithContext_posString___closed__2; +x_35 = lean_string_append(x_34, x_33); +lean_dec(x_33); +x_36 = lean_string_append(x_35, x_20); +x_37 = lean_ctor_get(x_30, 1); +lean_inc(x_37); +lean_dec(x_30); +x_38 = l_Nat_repr(x_37); +x_39 = lean_string_append(x_36, x_38); +lean_dec(x_38); +x_40 = lean_string_append(x_39, x_18); +x_41 = lean_string_append(x_25, x_40); +lean_dec(x_40); +x_42 = lean_string_append(x_41, 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_6); +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_dec(x_31); +x_44 = lean_ctor_get(x_30, 1); +lean_inc(x_44); +lean_dec(x_30); +x_45 = l_Nat_repr(x_44); +x_46 = l_Lean_Elab_WF_GuessLex_RecCallWithContext_posString___closed__2; +x_47 = lean_string_append(x_46, x_45); +lean_dec(x_45); +x_48 = lean_string_append(x_47, x_18); +x_49 = lean_string_append(x_25, x_48); +lean_dec(x_48); +x_50 = lean_string_append(x_49, x_18); +x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_6); +return x_51; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_RecCallWithContext_posString___boxed(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_WF_GuessLex_RecCallWithContext_posString(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_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_explainNonMutualFailure___spec__1(size_t x_1, size_t x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) +{ +return x_3; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; lean_object* x_10; size_t x_11; size_t x_12; lean_object* x_13; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = lean_erase_macro_scopes(x_5); +x_9 = 1; +x_10 = l_Lean_Name_toString(x_8, x_9); +x_11 = 1; +x_12 = lean_usize_add(x_2, x_11); +x_13 = lean_array_uset(x_7, x_2, x_10); +x_2 = x_12; +x_3 = x_13; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainNonMutualFailure___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +uint8_t x_12; +x_12 = lean_nat_dec_le(x_4, x_3); +if (x_12 == 0) +{ +lean_object* x_13; uint8_t x_14; +x_13 = lean_unsigned_to_nat(0u); +x_14 = lean_nat_dec_eq(x_2, x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_15 = lean_unsigned_to_nat(1u); +x_16 = lean_nat_sub(x_2, x_15); +lean_dec(x_2); +x_17 = l_Lean_Elab_WF_GuessLex_RecCallCache_prettyEntry(x_1, x_3, x_3, x_7, x_8, x_9, x_10, x_11); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_array_push(x_6, x_18); +x_21 = lean_nat_add(x_3, x_5); +lean_dec(x_3); +x_2 = x_16; +x_3 = x_21; +x_6 = x_20; +x_11 = x_19; +goto _start; +} +else +{ +lean_object* x_23; +lean_dec(x_3); +lean_dec(x_2); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_6); +lean_ctor_set(x_23, 1, x_11); +return x_23; +} +} +else +{ +lean_object* x_24; +lean_dec(x_3); +lean_dec(x_2); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_6); +lean_ctor_set(x_24, 1, x_11); +return x_24; +} +} +} +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainNonMutualFailure___spec__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(") ", 2); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainNonMutualFailure___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; +x_13 = lean_nat_dec_le(x_5, x_4); +if (x_13 == 0) +{ +lean_object* x_14; uint8_t x_15; +x_14 = lean_unsigned_to_nat(0u); +x_15 = lean_nat_dec_eq(x_3, x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_16 = lean_unsigned_to_nat(1u); +x_17 = lean_nat_sub(x_3, x_16); +lean_dec(x_3); +x_18 = !lean_is_exclusive(x_7); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_19 = lean_ctor_get(x_7, 0); +x_20 = lean_ctor_get(x_7, 1); +x_21 = lean_ctor_get(x_19, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_19, 1); +lean_inc(x_22); +x_23 = lean_ctor_get(x_19, 2); +lean_inc(x_23); +x_24 = lean_nat_dec_lt(x_22, x_23); +if (x_24 == 0) +{ +lean_object* x_25; +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_17); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_7); +lean_ctor_set(x_25, 1, x_12); +return x_25; +} +else +{ +uint8_t x_26; +x_26 = !lean_is_exclusive(x_19); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_27 = lean_ctor_get(x_19, 2); +lean_dec(x_27); +x_28 = lean_ctor_get(x_19, 1); +lean_dec(x_28); +x_29 = lean_ctor_get(x_19, 0); +lean_dec(x_29); +x_30 = lean_array_fget(x_21, x_22); +x_31 = lean_nat_add(x_22, x_16); +lean_dec(x_22); +lean_ctor_set(x_19, 1, x_31); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +x_33 = l_Lean_Elab_WF_GuessLex_RecCallWithContext_posString(x_32, x_8, x_9, x_10, 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 = lean_nat_add(x_4, x_16); +x_37 = l_Nat_repr(x_36); +x_38 = l_Lean_Elab_WF_GuessLex_initFn____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_9____closed__3; +x_39 = lean_string_append(x_38, x_37); +lean_dec(x_37); +x_40 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainNonMutualFailure___spec__3___closed__1; +x_41 = lean_string_append(x_39, x_40); +x_42 = lean_string_append(x_41, x_34); +lean_dec(x_34); +x_43 = lean_string_append(x_42, x_38); +lean_inc(x_2); +x_44 = lean_array_push(x_2, x_43); +lean_inc(x_1); +x_45 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainNonMutualFailure___spec__2(x_30, x_1, x_14, x_1, x_16, x_44, x_8, x_9, x_10, x_11, x_35); +lean_dec(x_30); +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_array_push(x_20, x_46); +lean_ctor_set(x_7, 1, x_48); +x_49 = lean_nat_add(x_4, x_6); +lean_dec(x_4); +x_3 = x_17; +x_4 = x_49; +x_12 = x_47; +goto _start; +} +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; 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_dec(x_19); +x_51 = lean_array_fget(x_21, x_22); +x_52 = lean_nat_add(x_22, x_16); +lean_dec(x_22); +x_53 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_53, 0, x_21); +lean_ctor_set(x_53, 1, x_52); +lean_ctor_set(x_53, 2, x_23); +x_54 = lean_ctor_get(x_51, 1); +lean_inc(x_54); +x_55 = l_Lean_Elab_WF_GuessLex_RecCallWithContext_posString(x_54, x_8, x_9, x_10, x_11, x_12); +x_56 = lean_ctor_get(x_55, 0); +lean_inc(x_56); +x_57 = lean_ctor_get(x_55, 1); +lean_inc(x_57); +lean_dec(x_55); +x_58 = lean_nat_add(x_4, x_16); +x_59 = l_Nat_repr(x_58); +x_60 = l_Lean_Elab_WF_GuessLex_initFn____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_9____closed__3; +x_61 = lean_string_append(x_60, x_59); +lean_dec(x_59); +x_62 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainNonMutualFailure___spec__3___closed__1; +x_63 = lean_string_append(x_61, x_62); +x_64 = lean_string_append(x_63, x_56); +lean_dec(x_56); +x_65 = lean_string_append(x_64, x_60); +lean_inc(x_2); +x_66 = lean_array_push(x_2, x_65); +lean_inc(x_1); +x_67 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainNonMutualFailure___spec__2(x_51, x_1, x_14, x_1, x_16, x_66, x_8, x_9, x_10, x_11, x_57); +lean_dec(x_51); +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_array_push(x_20, x_68); +lean_ctor_set(x_7, 1, x_70); +lean_ctor_set(x_7, 0, x_53); +x_71 = lean_nat_add(x_4, x_6); +lean_dec(x_4); +x_3 = x_17; +x_4 = x_71; +x_12 = x_69; +goto _start; +} +} +} +else +{ +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_73 = lean_ctor_get(x_7, 0); +x_74 = lean_ctor_get(x_7, 1); +lean_inc(x_74); +lean_inc(x_73); +lean_dec(x_7); +x_75 = lean_ctor_get(x_73, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_73, 1); +lean_inc(x_76); +x_77 = lean_ctor_get(x_73, 2); +lean_inc(x_77); +x_78 = lean_nat_dec_lt(x_76, x_77); +if (x_78 == 0) +{ +lean_object* x_79; lean_object* x_80; +lean_dec(x_77); +lean_dec(x_76); +lean_dec(x_75); +lean_dec(x_17); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_79 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_79, 0, x_73); +lean_ctor_set(x_79, 1, x_74); +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_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; +if (lean_is_exclusive(x_73)) { + lean_ctor_release(x_73, 0); + lean_ctor_release(x_73, 1); + lean_ctor_release(x_73, 2); + x_81 = x_73; +} else { + lean_dec_ref(x_73); + x_81 = lean_box(0); +} +x_82 = lean_array_fget(x_75, x_76); +x_83 = lean_nat_add(x_76, x_16); +lean_dec(x_76); +if (lean_is_scalar(x_81)) { + x_84 = lean_alloc_ctor(0, 3, 0); +} else { + x_84 = x_81; +} +lean_ctor_set(x_84, 0, x_75); +lean_ctor_set(x_84, 1, x_83); +lean_ctor_set(x_84, 2, x_77); +x_85 = lean_ctor_get(x_82, 1); +lean_inc(x_85); +x_86 = l_Lean_Elab_WF_GuessLex_RecCallWithContext_posString(x_85, x_8, x_9, x_10, x_11, x_12); +x_87 = lean_ctor_get(x_86, 0); +lean_inc(x_87); +x_88 = lean_ctor_get(x_86, 1); +lean_inc(x_88); +lean_dec(x_86); +x_89 = lean_nat_add(x_4, x_16); +x_90 = l_Nat_repr(x_89); +x_91 = l_Lean_Elab_WF_GuessLex_initFn____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_9____closed__3; +x_92 = lean_string_append(x_91, x_90); +lean_dec(x_90); +x_93 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainNonMutualFailure___spec__3___closed__1; +x_94 = lean_string_append(x_92, x_93); +x_95 = lean_string_append(x_94, x_87); +lean_dec(x_87); +x_96 = lean_string_append(x_95, x_91); +lean_inc(x_2); +x_97 = lean_array_push(x_2, x_96); +lean_inc(x_1); +x_98 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainNonMutualFailure___spec__2(x_82, x_1, x_14, x_1, x_16, x_97, x_8, x_9, x_10, x_11, x_88); +lean_dec(x_82); +x_99 = lean_ctor_get(x_98, 0); +lean_inc(x_99); +x_100 = lean_ctor_get(x_98, 1); +lean_inc(x_100); +lean_dec(x_98); +x_101 = lean_array_push(x_74, x_99); +x_102 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_102, 0, x_84); +lean_ctor_set(x_102, 1, x_101); +x_103 = lean_nat_add(x_4, x_6); +lean_dec(x_4); +x_3 = x_17; +x_4 = x_103; +x_7 = x_102; +x_12 = x_100; +goto _start; +} +} +} +else +{ +lean_object* x_105; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_105 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_105, 0, x_7); +lean_ctor_set(x_105, 1, x_12); +return x_105; +} +} +else +{ +lean_object* x_106; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_106 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_106, 0, x_7); +lean_ctor_set(x_106, 1, x_12); +return x_106; +} +} +} +static lean_object* _init_l_Lean_Elab_WF_GuessLex_explainNonMutualFailure___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_WF_GuessLex_GuessLexRel_toNatRel___closed__20; +x_2 = l_Lean_Elab_WF_GuessLex_initFn____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_9____closed__3; +x_3 = lean_array_push(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_explainNonMutualFailure(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_8 = lean_array_get_size(x_1); +x_9 = lean_usize_of_nat(x_8); +x_10 = 0; +x_11 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_explainNonMutualFailure___spec__1(x_9, x_10, x_1); +x_12 = l_Lean_Elab_WF_GuessLex_explainNonMutualFailure___closed__1; +x_13 = l_Array_append___rarg(x_12, x_11); +x_14 = l_Lean_Elab_WF_GuessLex_GuessLexRel_toNatRel___closed__20; +x_15 = lean_array_push(x_14, x_13); +x_16 = lean_array_get_size(x_2); +x_17 = lean_unsigned_to_nat(0u); +lean_inc(x_16); +x_18 = l_Array_toSubarray___rarg(x_2, x_17, 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_15); +x_20 = lean_unsigned_to_nat(1u); +lean_inc(x_16); +x_21 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainNonMutualFailure___spec__3(x_8, x_14, x_16, x_17, x_16, x_20, x_19, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_16); +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; +x_23 = lean_ctor_get(x_21, 0); +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +lean_dec(x_23); +x_25 = l_Lean_Elab_WF_GuessLex_formatTable(x_24); +lean_dec(x_24); +x_26 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_26, 0, x_25); +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_object* x_30; lean_object* x_31; lean_object* x_32; +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_ctor_get(x_27, 1); +lean_inc(x_29); +lean_dec(x_27); +x_30 = l_Lean_Elab_WF_GuessLex_formatTable(x_29); +lean_dec(x_29); +x_31 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_31, 0, 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; +} +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_explainNonMutualFailure___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_explainNonMutualFailure___spec__1(x_4, x_5, x_3); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainNonMutualFailure___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) { +_start: +{ +lean_object* x_12; +x_12 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainNonMutualFailure___spec__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_5); +lean_dec(x_4); +lean_dec(x_1); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainNonMutualFailure___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_object* x_13; +x_13 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainNonMutualFailure___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_explainNonMutualFailure___boxed(lean_object* x_1, 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_WF_GuessLex_explainNonMutualFailure(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_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; +x_13 = lean_nat_dec_le(x_5, x_4); +if (x_13 == 0) +{ +lean_object* x_14; uint8_t x_15; +x_14 = lean_unsigned_to_nat(0u); +x_15 = lean_nat_dec_eq(x_3, x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_16 = lean_unsigned_to_nat(1u); +x_17 = lean_nat_sub(x_3, x_16); +lean_dec(x_3); +x_18 = l_Lean_Elab_WF_GuessLex_RecCallCache_prettyEntry(x_1, x_4, x_2, x_8, x_9, x_10, x_11, x_12); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_array_push(x_7, x_19); +x_22 = lean_nat_add(x_4, x_6); +lean_dec(x_4); +x_3 = x_17; +x_4 = x_22; +x_7 = x_21; +x_12 = x_20; +goto _start; +} +else +{ +lean_object* x_24; +lean_dec(x_4); +lean_dec(x_3); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_7); +lean_ctor_set(x_24, 1, x_12); +return x_24; +} +} +else +{ +lean_object* x_25; +lean_dec(x_4); +lean_dec(x_3); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_7); +lean_ctor_set(x_25, 1, x_12); +return x_25; +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; +x_13 = lean_nat_dec_le(x_5, x_4); +if (x_13 == 0) +{ +lean_object* x_14; uint8_t x_15; +x_14 = lean_unsigned_to_nat(0u); +x_15 = lean_nat_dec_eq(x_3, x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_16 = lean_unsigned_to_nat(1u); +x_17 = lean_nat_sub(x_3, x_16); +lean_dec(x_3); +x_18 = l_Lean_Elab_WF_GuessLex_RecCallCache_prettyEntry(x_1, x_4, x_2, x_8, x_9, x_10, x_11, x_12); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_array_push(x_7, x_19); +x_22 = lean_nat_add(x_4, x_6); +lean_dec(x_4); +x_3 = x_17; +x_4 = x_22; +x_7 = x_21; +x_12 = x_20; +goto _start; +} +else +{ +lean_object* x_24; +lean_dec(x_4); +lean_dec(x_3); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_7); +lean_ctor_set(x_24, 1, x_12); +return x_24; +} +} +else +{ +lean_object* x_25; +lean_dec(x_4); +lean_dec(x_3); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_7); +lean_ctor_set(x_25, 1, x_12); +return x_25; +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +uint8_t x_15; +x_15 = lean_nat_dec_le(x_7, x_6); +if (x_15 == 0) +{ +lean_object* x_16; uint8_t x_17; +x_16 = lean_unsigned_to_nat(0u); +x_17 = lean_nat_dec_eq(x_5, x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_18 = lean_unsigned_to_nat(1u); +x_19 = lean_nat_sub(x_5, x_18); +lean_dec(x_5); +x_20 = lean_nat_dec_lt(x_6, x_4); +if (x_20 == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_27 = lean_ctor_get(x_19, 2); -lean_dec(x_27); -x_28 = lean_ctor_get(x_19, 1); -lean_dec(x_28); -x_29 = lean_ctor_get(x_19, 0); -lean_dec(x_29); -x_30 = lean_array_fget(x_21, x_22); -x_31 = lean_nat_add(x_22, x_16); -lean_dec(x_22); -lean_ctor_set(x_19, 1, x_31); -x_32 = lean_ctor_get(x_30, 1); -lean_inc(x_32); -x_33 = l_Lean_Elab_WF_GuessLex_RecCallWithContext_posString(x_32, x_8, x_9, x_10, 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 = lean_nat_add(x_4, x_16); -x_37 = l_Nat_repr(x_36); -x_38 = l_Lean_Elab_WF_GuessLex_initFn____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_9____closed__3; -x_39 = lean_string_append(x_38, x_37); -lean_dec(x_37); -x_40 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainNonMutualFailure___spec__3___closed__1; -x_41 = lean_string_append(x_39, x_40); -x_42 = lean_string_append(x_41, x_34); -lean_dec(x_34); -x_43 = lean_string_append(x_42, x_38); +lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_21 = l_Lean_instInhabitedName; +x_22 = l___private_Init_Util_0__outOfBounds___rarg(x_21); +x_23 = lean_erase_macro_scopes(x_22); +x_24 = 1; +x_25 = l_Lean_Name_toString(x_23, x_24); +x_26 = l_Lean_Elab_WF_GuessLex_naryVarNames___closed__1; +x_27 = lean_array_push(x_26, x_25); lean_inc(x_2); -x_44 = lean_array_push(x_2, x_43); -lean_inc(x_1); -x_45 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainNonMutualFailure___spec__2(x_30, x_1, x_14, x_1, x_16, x_44, x_8, x_9, x_10, x_11, x_35); -lean_dec(x_30); -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_array_push(x_20, x_46); -lean_ctor_set(x_7, 1, x_48); -x_49 = lean_nat_add(x_4, x_6); -lean_dec(x_4); -x_3 = x_17; -x_4 = x_49; -x_12 = x_47; +x_28 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__1(x_1, x_6, x_2, x_16, x_2, x_18, x_27, x_10, x_11, x_12, x_13, x_14); +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_array_push(x_9, x_29); +x_32 = lean_nat_add(x_6, x_8); +lean_dec(x_6); +x_5 = x_19; +x_6 = x_32; +x_9 = x_31; +x_14 = x_30; goto _start; } 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; 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_dec(x_19); -x_51 = lean_array_fget(x_21, x_22); -x_52 = lean_nat_add(x_22, x_16); -lean_dec(x_22); -x_53 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_53, 0, x_21); -lean_ctor_set(x_53, 1, x_52); -lean_ctor_set(x_53, 2, x_23); -x_54 = lean_ctor_get(x_51, 1); -lean_inc(x_54); -x_55 = l_Lean_Elab_WF_GuessLex_RecCallWithContext_posString(x_54, x_8, x_9, x_10, x_11, x_12); -x_56 = lean_ctor_get(x_55, 0); -lean_inc(x_56); -x_57 = lean_ctor_get(x_55, 1); -lean_inc(x_57); -lean_dec(x_55); -x_58 = lean_nat_add(x_4, x_16); -x_59 = l_Nat_repr(x_58); -x_60 = l_Lean_Elab_WF_GuessLex_initFn____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_9____closed__3; -x_61 = lean_string_append(x_60, x_59); -lean_dec(x_59); -x_62 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainNonMutualFailure___spec__3___closed__1; -x_63 = lean_string_append(x_61, x_62); -x_64 = lean_string_append(x_63, x_56); -lean_dec(x_56); -x_65 = lean_string_append(x_64, x_60); +lean_object* x_34; lean_object* x_35; uint8_t x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_34 = lean_array_fget(x_3, x_6); +x_35 = lean_erase_macro_scopes(x_34); +x_36 = 1; +x_37 = l_Lean_Name_toString(x_35, x_36); +x_38 = l_Lean_Elab_WF_GuessLex_naryVarNames___closed__1; +x_39 = lean_array_push(x_38, x_37); lean_inc(x_2); -x_66 = lean_array_push(x_2, x_65); -lean_inc(x_1); -x_67 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainNonMutualFailure___spec__2(x_51, x_1, x_14, x_1, x_16, x_66, x_8, x_9, x_10, x_11, x_57); -lean_dec(x_51); -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_array_push(x_20, x_68); -lean_ctor_set(x_7, 1, x_70); -lean_ctor_set(x_7, 0, x_53); -x_71 = lean_nat_add(x_4, x_6); -lean_dec(x_4); -x_3 = x_17; -x_4 = x_71; -x_12 = x_69; +x_40 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__2(x_1, x_6, x_2, x_16, x_2, x_18, x_39, x_10, x_11, x_12, x_13, x_14); +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_array_push(x_9, x_41); +x_44 = lean_nat_add(x_6, x_8); +lean_dec(x_6); +x_5 = x_19; +x_6 = x_44; +x_9 = x_43; +x_14 = x_42; goto _start; } } -} else { -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_73 = lean_ctor_get(x_7, 0); -x_74 = lean_ctor_get(x_7, 1); -lean_inc(x_74); -lean_inc(x_73); -lean_dec(x_7); -x_75 = lean_ctor_get(x_73, 0); -lean_inc(x_75); -x_76 = lean_ctor_get(x_73, 1); -lean_inc(x_76); -x_77 = lean_ctor_get(x_73, 2); -lean_inc(x_77); -x_78 = lean_nat_dec_lt(x_76, x_77); -if (x_78 == 0) -{ -lean_object* x_79; lean_object* x_80; -lean_dec(x_77); -lean_dec(x_76); -lean_dec(x_75); -lean_dec(x_17); -lean_dec(x_4); +lean_object* x_46; +lean_dec(x_6); +lean_dec(x_5); lean_dec(x_2); -lean_dec(x_1); -x_79 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_79, 0, x_73); -lean_ctor_set(x_79, 1, x_74); -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; +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_9); +lean_ctor_set(x_46, 1, x_14); +return x_46; +} } else { -lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_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; -if (lean_is_exclusive(x_73)) { - lean_ctor_release(x_73, 0); - lean_ctor_release(x_73, 1); - lean_ctor_release(x_73, 2); - x_81 = x_73; -} else { - lean_dec_ref(x_73); - x_81 = lean_box(0); +lean_object* x_47; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_9); +lean_ctor_set(x_47, 1, x_14); +return x_47; } -x_82 = lean_array_fget(x_75, x_76); -x_83 = lean_nat_add(x_76, x_16); -lean_dec(x_76); -if (lean_is_scalar(x_81)) { - x_84 = lean_alloc_ctor(0, 3, 0); -} else { - x_84 = x_81; } -lean_ctor_set(x_84, 0, x_75); -lean_ctor_set(x_84, 1, x_83); -lean_ctor_set(x_84, 2, x_77); -x_85 = lean_ctor_get(x_82, 1); -lean_inc(x_85); -x_86 = l_Lean_Elab_WF_GuessLex_RecCallWithContext_posString(x_85, x_8, x_9, x_10, x_11, x_12); -x_87 = lean_ctor_get(x_86, 0); -lean_inc(x_87); -x_88 = lean_ctor_get(x_86, 1); -lean_inc(x_88); -lean_dec(x_86); -x_89 = lean_nat_add(x_4, x_16); -x_90 = l_Nat_repr(x_89); -x_91 = l_Lean_Elab_WF_GuessLex_initFn____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_9____closed__3; -x_92 = lean_string_append(x_91, x_90); -lean_dec(x_90); -x_93 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainNonMutualFailure___spec__3___closed__1; -x_94 = lean_string_append(x_92, x_93); -x_95 = lean_string_append(x_94, x_87); -lean_dec(x_87); -x_96 = lean_string_append(x_95, x_91); -lean_inc(x_2); -x_97 = lean_array_push(x_2, x_96); -lean_inc(x_1); -x_98 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainNonMutualFailure___spec__2(x_82, x_1, x_14, x_1, x_16, x_97, x_8, x_9, x_10, x_11, x_88); -lean_dec(x_82); -x_99 = lean_ctor_get(x_98, 0); -lean_inc(x_99); -x_100 = lean_ctor_get(x_98, 1); -lean_inc(x_100); -lean_dec(x_98); -x_101 = lean_array_push(x_74, x_99); -x_102 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_102, 0, x_84); -lean_ctor_set(x_102, 1, x_101); -x_103 = lean_nat_add(x_4, x_6); -lean_dec(x_4); -x_3 = x_17; -x_4 = x_103; -x_7 = x_102; -x_12 = x_100; +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +uint8_t x_12; +x_12 = lean_nat_dec_le(x_4, x_3); +if (x_12 == 0) +{ +lean_object* x_13; uint8_t x_14; +x_13 = lean_unsigned_to_nat(0u); +x_14 = lean_nat_dec_eq(x_2, x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_15 = lean_unsigned_to_nat(1u); +x_16 = lean_nat_sub(x_2, x_15); +lean_dec(x_2); +x_17 = l_Lean_Elab_WF_GuessLex_RecCallCache_prettyEntry(x_1, x_3, x_3, x_7, x_8, x_9, x_10, x_11); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_array_push(x_6, x_18); +x_21 = lean_nat_add(x_3, x_5); +lean_dec(x_3); +x_2 = x_16; +x_3 = x_21; +x_6 = x_20; +x_11 = x_19; goto _start; } -} -} else { -lean_object* x_105; -lean_dec(x_4); +lean_object* x_23; lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_105 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_105, 0, x_7); -lean_ctor_set(x_105, 1, x_12); -return x_105; +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_6); +lean_ctor_set(x_23, 1, x_11); +return x_23; } } else { -lean_object* x_106; -lean_dec(x_4); +lean_object* x_24; lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_106 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_106, 0, x_7); -lean_ctor_set(x_106, 1, x_12); -return x_106; +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_6); +lean_ctor_set(x_24, 1, x_11); +return x_24; } } } -static lean_object* _init_l_Lean_Elab_WF_GuessLex_explainNonMutualFailure___closed__1() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___lambda__1___closed__1() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_WF_GuessLex_GuessLexRel_toNatRel___closed__20; -x_2 = l_Lean_Elab_WF_GuessLex_initFn____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_9____closed__3; -x_3 = lean_array_push(x_1, x_2); -return x_3; +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__6___closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_explainNonMutualFailure(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_8 = lean_array_get_size(x_1); -x_9 = lean_usize_of_nat(x_8); -x_10 = 0; -x_11 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_explainNonMutualFailure___spec__1(x_9, x_10, x_1); -x_12 = l_Lean_Elab_WF_GuessLex_explainNonMutualFailure___closed__1; -x_13 = l_Array_append___rarg(x_12, x_11); -x_14 = l_Lean_Elab_WF_GuessLex_GuessLexRel_toNatRel___closed__20; -x_15 = lean_array_push(x_14, x_13); -x_16 = lean_array_get_size(x_2); -x_17 = lean_unsigned_to_nat(0u); -lean_inc(x_16); -x_18 = l_Array_toSubarray___rarg(x_2, x_17, 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_15); -x_20 = lean_unsigned_to_nat(1u); -lean_inc(x_16); -x_21 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainNonMutualFailure___spec__3(x_8, x_14, x_16, x_17, x_16, x_20, x_19, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_16); -x_22 = !lean_is_exclusive(x_21); -if (x_22 == 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; +x_9 = l_Lean_Elab_WF_GuessLex_formatTable(x_2); +x_10 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_10, 0, x_9); +x_11 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_11, 0, x_1); +lean_ctor_set(x_11, 1, x_10); +x_12 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___lambda__1___closed__1; +x_13 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_13, 0, x_11); +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_8); +return x_15; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___closed__1() { +_start: { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_23 = lean_ctor_get(x_21, 0); -x_24 = lean_ctor_get(x_23, 1); -lean_inc(x_24); -lean_dec(x_23); -x_25 = l_Lean_Elab_WF_GuessLex_formatTable(x_24); -lean_dec(x_24); -x_26 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_21, 0, x_26); -return x_21; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Call from ", 10); +return x_1; } -else +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___closed__2() { +_start: { -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_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_ctor_get(x_27, 1); -lean_inc(x_29); -lean_dec(x_27); -x_30 = l_Lean_Elab_WF_GuessLex_formatTable(x_29); -lean_dec(x_29); -x_31 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_31, 0, 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; +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___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_Elab_WF_GuessLex_explainMutualFailure___spec__5___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(" to ", 4); +return x_1; +} } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_explainNonMutualFailure___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___closed__4() { _start: { -size_t x_4; size_t x_5; lean_object* x_6; -x_4 = lean_unbox_usize(x_1); -lean_dec(x_1); -x_5 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_explainNonMutualFailure___spec__1(x_4, x_5, x_3); -return x_6; +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___closed__3; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainNonMutualFailure___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) { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___closed__5() { _start: { -lean_object* x_12; -x_12 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainNonMutualFailure___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__4___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_Elab_WF_GuessLex_explainMutualFailure___spec__5___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("at ", 3); +return x_1; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___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_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(":\n", 2); +return x_1; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___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_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___lambda__1___boxed), 8, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___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_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +uint8_t x_12; +x_12 = lean_usize_dec_lt(x_5, x_4); +if (x_12 == 0) +{ +lean_object* x_13; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -return x_12; +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_6); +lean_ctor_set(x_13, 1, x_11); +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; uint8_t 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; uint8_t x_30; size_t x_31; lean_object* x_32; uint8_t x_33; lean_object* x_34; +x_14 = lean_array_uget(x_3, x_5); +x_15 = lean_ctor_get(x_14, 1); +lean_inc(x_15); +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 3); +lean_inc(x_17); +x_18 = l_Lean_Elab_WF_GuessLex_RecCallWithContext_posString(x_15, 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_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_array_get_size(x_1); +x_22 = lean_nat_dec_lt(x_16, x_21); +x_23 = lean_nat_dec_lt(x_17, x_21); +lean_dec(x_21); +x_24 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_24, 0, x_19); +x_25 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___closed__7; +x_26 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_24); +x_27 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___closed__9; +x_28 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +x_29 = lean_array_get_size(x_2); +x_30 = lean_nat_dec_lt(x_16, x_29); +x_31 = 0; +x_32 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___closed__10; +x_33 = lean_nat_dec_eq(x_16, x_17); +if (x_22 == 0) +{ +lean_object* x_121; lean_object* x_122; +x_121 = l_Lean_instInhabitedName; +x_122 = l___private_Init_Util_0__outOfBounds___rarg(x_121); +x_34 = x_122; +goto block_120; } +else +{ +lean_object* x_123; +x_123 = lean_array_fget(x_1, x_16); +x_34 = x_123; +goto block_120; } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainNonMutualFailure___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: +block_120: { -lean_object* x_13; -x_13 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainNonMutualFailure___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -return x_13; +uint8_t x_35; lean_object* x_36; 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_35 = 1; +x_36 = l_Lean_Name_toString(x_34, x_35); +x_37 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_37, 0, x_36); +x_38 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___closed__2; +x_39 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_37); +x_40 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___closed__4; +x_41 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +if (x_23 == 0) +{ +lean_object* x_117; lean_object* x_118; +x_117 = l_Lean_instInhabitedName; +x_118 = l___private_Init_Util_0__outOfBounds___rarg(x_117); +x_42 = x_118; +goto block_116; } +else +{ +lean_object* x_119; +x_119 = lean_array_fget(x_1, x_17); +x_42 = x_119; +goto block_116; } -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_explainNonMutualFailure___boxed(lean_object* x_1, 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: +block_116: { -lean_object* x_8; -x_8 = l_Lean_Elab_WF_GuessLex_explainNonMutualFailure(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_43; lean_object* x_44; 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_43 = l_Lean_Name_toString(x_42, x_35); +x_44 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_44, 0, x_43); +x_45 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_45, 0, x_41); +lean_ctor_set(x_45, 1, x_44); +x_46 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___closed__5; +x_47 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +x_48 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_48, 0, x_6); +lean_ctor_set(x_48, 1, x_47); +x_49 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_28); +if (x_30 == 0) +{ +lean_object* x_113; lean_object* x_114; +lean_dec(x_16); +x_113 = l_Lean_Elab_WF_GuessLex_naryVarNames___closed__1; +x_114 = l___private_Init_Util_0__outOfBounds___rarg(x_113); +x_50 = x_114; +goto block_112; } +else +{ +lean_object* x_115; +x_115 = lean_array_fget(x_2, x_16); +lean_dec(x_16); +x_50 = x_115; +goto block_112; } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: +block_112: { -uint8_t x_13; -x_13 = lean_nat_dec_le(x_5, x_4); -if (x_13 == 0) +lean_object* x_51; size_t x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_51 = lean_array_get_size(x_50); +x_52 = lean_usize_of_nat(x_51); +x_53 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_explainNonMutualFailure___spec__1(x_52, x_31, x_50); +x_54 = l_Lean_Elab_WF_GuessLex_explainNonMutualFailure___closed__1; +x_55 = l_Array_append___rarg(x_54, x_53); +x_56 = l_Lean_Elab_WF_GuessLex_GuessLexRel_toNatRel___closed__20; +x_57 = lean_array_push(x_56, x_55); +if (x_33 == 0) { -lean_object* x_14; uint8_t x_15; -x_14 = lean_unsigned_to_nat(0u); -x_15 = lean_nat_dec_eq(x_3, x_14); -if (x_15 == 0) +uint8_t x_84; +x_84 = lean_nat_dec_lt(x_17, x_29); +lean_dec(x_29); +if (x_84 == 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; -x_16 = lean_unsigned_to_nat(1u); -x_17 = lean_nat_sub(x_3, x_16); -lean_dec(x_3); -x_18 = l_Lean_Elab_WF_GuessLex_RecCallCache_prettyEntry(x_1, x_4, x_2, x_8, x_9, x_10, x_11, x_12); -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_18, 1); -lean_inc(x_20); -lean_dec(x_18); -x_21 = lean_array_push(x_7, x_19); -x_22 = lean_nat_add(x_4, x_6); -lean_dec(x_4); -x_3 = x_17; -x_4 = x_22; -x_7 = x_21; -x_12 = x_20; -goto _start; +lean_object* x_85; lean_object* x_86; +lean_dec(x_17); +x_85 = l_Lean_Elab_WF_GuessLex_naryVarNames___closed__1; +x_86 = l___private_Init_Util_0__outOfBounds___rarg(x_85); +x_58 = x_86; +goto block_83; } else { -lean_object* x_24; -lean_dec(x_4); -lean_dec(x_3); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_7); -lean_ctor_set(x_24, 1, x_12); -return x_24; +lean_object* x_87; +x_87 = lean_array_fget(x_2, x_17); +lean_dec(x_17); +x_58 = x_87; +goto block_83; } } else { -lean_object* x_25; -lean_dec(x_4); -lean_dec(x_3); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_7); -lean_ctor_set(x_25, 1, x_12); -return x_25; +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; +lean_dec(x_29); +lean_dec(x_17); +x_88 = lean_unsigned_to_nat(0u); +x_89 = lean_unsigned_to_nat(1u); +lean_inc(x_51); +x_90 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__4(x_14, x_51, x_88, x_51, x_89, x_54, x_7, x_8, x_9, x_10, x_20); +lean_dec(x_51); +lean_dec(x_14); +x_91 = lean_ctor_get(x_90, 0); +lean_inc(x_91); +x_92 = lean_ctor_get(x_90, 1); +lean_inc(x_92); +lean_dec(x_90); +x_93 = lean_array_push(x_57, x_91); +x_94 = lean_box(0); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_95 = lean_apply_8(x_32, x_49, x_93, x_94, x_7, x_8, x_9, x_10, x_92); +if (lean_obj_tag(x_95) == 0) +{ +lean_object* x_96; +x_96 = lean_ctor_get(x_95, 0); +lean_inc(x_96); +if (lean_obj_tag(x_96) == 0) +{ +uint8_t x_97; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +x_97 = !lean_is_exclusive(x_95); +if (x_97 == 0) +{ +lean_object* x_98; lean_object* x_99; +x_98 = lean_ctor_get(x_95, 0); +lean_dec(x_98); +x_99 = lean_ctor_get(x_96, 0); +lean_inc(x_99); +lean_dec(x_96); +lean_ctor_set(x_95, 0, x_99); +return x_95; } +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_ctor_get(x_96, 0); +lean_inc(x_101); +lean_dec(x_96); +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_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -uint8_t x_13; -x_13 = lean_nat_dec_le(x_5, x_4); -if (x_13 == 0) -{ -lean_object* x_14; uint8_t x_15; -x_14 = lean_unsigned_to_nat(0u); -x_15 = lean_nat_dec_eq(x_3, x_14); -if (x_15 == 0) +else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_16 = lean_unsigned_to_nat(1u); -x_17 = lean_nat_sub(x_3, x_16); -lean_dec(x_3); -x_18 = l_Lean_Elab_WF_GuessLex_RecCallCache_prettyEntry(x_1, x_4, x_2, x_8, x_9, x_10, x_11, x_12); -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_18, 1); -lean_inc(x_20); -lean_dec(x_18); -x_21 = lean_array_push(x_7, x_19); -x_22 = lean_nat_add(x_4, x_6); -lean_dec(x_4); -x_3 = x_17; -x_4 = x_22; -x_7 = x_21; -x_12 = x_20; +lean_object* x_103; lean_object* x_104; size_t x_105; size_t x_106; +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 = 1; +x_106 = lean_usize_add(x_5, x_105); +x_5 = x_106; +x_6 = x_104; +x_11 = x_103; goto _start; } +} else { -lean_object* x_24; -lean_dec(x_4); -lean_dec(x_3); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_7); -lean_ctor_set(x_24, 1, x_12); -return x_24; -} +uint8_t x_108; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +x_108 = !lean_is_exclusive(x_95); +if (x_108 == 0) +{ +return x_95; } else { -lean_object* x_25; -lean_dec(x_4); -lean_dec(x_3); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_7); -lean_ctor_set(x_25, 1, x_12); -return x_25; +lean_object* x_109; lean_object* x_110; lean_object* x_111; +x_109 = lean_ctor_get(x_95, 0); +x_110 = lean_ctor_get(x_95, 1); +lean_inc(x_110); +lean_inc(x_109); +lean_dec(x_95); +x_111 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_111, 0, x_109); +lean_ctor_set(x_111, 1, x_110); +return x_111; } } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: +block_83: { -uint8_t x_15; -x_15 = lean_nat_dec_le(x_7, x_6); -if (x_15 == 0) +lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_59 = lean_array_get_size(x_58); +x_60 = lean_unsigned_to_nat(0u); +x_61 = lean_unsigned_to_nat(1u); +lean_inc(x_59); +x_62 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__3(x_14, x_51, x_58, x_59, x_59, x_60, x_59, x_61, x_57, x_7, x_8, x_9, x_10, x_20); +lean_dec(x_59); +lean_dec(x_58); +lean_dec(x_14); +x_63 = lean_ctor_get(x_62, 0); +lean_inc(x_63); +x_64 = lean_ctor_get(x_62, 1); +lean_inc(x_64); +lean_dec(x_62); +x_65 = lean_box(0); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_66 = lean_apply_8(x_32, x_49, x_63, x_65, x_7, x_8, x_9, x_10, x_64); +if (lean_obj_tag(x_66) == 0) { -lean_object* x_16; uint8_t x_17; -x_16 = lean_unsigned_to_nat(0u); -x_17 = lean_nat_dec_eq(x_5, x_16); -if (x_17 == 0) +lean_object* x_67; +x_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +if (lean_obj_tag(x_67) == 0) { -lean_object* x_18; lean_object* x_19; uint8_t x_20; -x_18 = lean_unsigned_to_nat(1u); -x_19 = lean_nat_sub(x_5, x_18); -lean_dec(x_5); -x_20 = lean_nat_dec_lt(x_6, x_4); -if (x_20 == 0) +uint8_t x_68; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +x_68 = !lean_is_exclusive(x_66); +if (x_68 == 0) { -lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_21 = l_Lean_instInhabitedName; -x_22 = l___private_Init_Util_0__outOfBounds___rarg(x_21); -x_23 = lean_erase_macro_scopes(x_22); -x_24 = 1; -x_25 = l_Lean_Name_toString(x_23, x_24); -x_26 = l_Lean_Elab_WF_GuessLex_naryVarNames___closed__1; -x_27 = lean_array_push(x_26, x_25); -lean_inc(x_2); -x_28 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__1(x_1, x_6, x_2, x_16, x_2, x_18, x_27, x_10, x_11, x_12, x_13, x_14); -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_array_push(x_9, x_29); -x_32 = lean_nat_add(x_6, x_8); -lean_dec(x_6); -x_5 = x_19; -x_6 = x_32; -x_9 = x_31; -x_14 = x_30; -goto _start; +lean_object* x_69; lean_object* x_70; +x_69 = lean_ctor_get(x_66, 0); +lean_dec(x_69); +x_70 = lean_ctor_get(x_67, 0); +lean_inc(x_70); +lean_dec(x_67); +lean_ctor_set(x_66, 0, x_70); +return x_66; } else { -lean_object* x_34; lean_object* x_35; uint8_t x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_34 = lean_array_fget(x_3, x_6); -x_35 = lean_erase_macro_scopes(x_34); -x_36 = 1; -x_37 = l_Lean_Name_toString(x_35, x_36); -x_38 = l_Lean_Elab_WF_GuessLex_naryVarNames___closed__1; -x_39 = lean_array_push(x_38, x_37); -lean_inc(x_2); -x_40 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__2(x_1, x_6, x_2, x_16, x_2, x_18, x_39, x_10, x_11, x_12, x_13, x_14); -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_array_push(x_9, x_41); -x_44 = lean_nat_add(x_6, x_8); -lean_dec(x_6); -x_5 = x_19; -x_6 = x_44; -x_9 = x_43; -x_14 = x_42; -goto _start; +lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_71 = lean_ctor_get(x_66, 1); +lean_inc(x_71); +lean_dec(x_66); +x_72 = lean_ctor_get(x_67, 0); +lean_inc(x_72); +lean_dec(x_67); +x_73 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_73, 0, x_72); +lean_ctor_set(x_73, 1, x_71); +return x_73; } } else { -lean_object* x_46; -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_2); -x_46 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_46, 0, x_9); -lean_ctor_set(x_46, 1, x_14); -return x_46; +lean_object* x_74; lean_object* x_75; size_t x_76; size_t x_77; +x_74 = lean_ctor_get(x_66, 1); +lean_inc(x_74); +lean_dec(x_66); +x_75 = lean_ctor_get(x_67, 0); +lean_inc(x_75); +lean_dec(x_67); +x_76 = 1; +x_77 = lean_usize_add(x_5, x_76); +x_5 = x_77; +x_6 = x_75; +x_11 = x_74; +goto _start; } } else { -lean_object* x_47; -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_2); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_9); -lean_ctor_set(x_47, 1, x_14); -return x_47; -} -} -} -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -uint8_t x_12; -x_12 = lean_nat_dec_le(x_4, x_3); -if (x_12 == 0) -{ -lean_object* x_13; uint8_t x_14; -x_13 = lean_unsigned_to_nat(0u); -x_14 = lean_nat_dec_eq(x_2, x_13); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_15 = lean_unsigned_to_nat(1u); -x_16 = lean_nat_sub(x_2, x_15); -lean_dec(x_2); -x_17 = l_Lean_Elab_WF_GuessLex_RecCallCache_prettyEntry(x_1, x_3, x_3, x_7, x_8, x_9, x_10, x_11); -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_17, 1); -lean_inc(x_19); -lean_dec(x_17); -x_20 = lean_array_push(x_6, x_18); -x_21 = lean_nat_add(x_3, x_5); -lean_dec(x_3); -x_2 = x_16; -x_3 = x_21; -x_6 = x_20; -x_11 = x_19; -goto _start; -} -else +uint8_t x_79; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +x_79 = !lean_is_exclusive(x_66); +if (x_79 == 0) { -lean_object* x_23; -lean_dec(x_3); -lean_dec(x_2); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_6); -lean_ctor_set(x_23, 1, x_11); -return x_23; -} +return x_66; } else { -lean_object* x_24; -lean_dec(x_3); -lean_dec(x_2); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_6); -lean_ctor_set(x_24, 1, x_11); -return x_24; -} +lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_80 = lean_ctor_get(x_66, 0); +x_81 = lean_ctor_get(x_66, 1); +lean_inc(x_81); +lean_inc(x_80); +lean_dec(x_66); +x_82 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_82, 0, x_80); +lean_ctor_set(x_82, 1, x_81); +return x_82; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___lambda__1___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__6___closed__1; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_9 = l_Lean_Elab_WF_GuessLex_formatTable(x_2); -x_10 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_10, 0, x_9); -x_11 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_11, 0, x_1); -lean_ctor_set(x_11, 1, x_10); -x_12 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___lambda__1___closed__1; -x_13 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_13, 0, x_11); -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_8); -return x_15; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Call from ", 10); -return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___closed__2() { +} +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_explainMutualFailure(lean_object* x_1, lean_object* x_2, 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_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___closed__1; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} +lean_object* x_9; size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; +x_9 = lean_array_get_size(x_3); +x_10 = lean_usize_of_nat(x_9); +lean_dec(x_9); +x_11 = 0; +x_12 = lean_box(0); +x_13 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5(x_1, x_2, x_3, x_10, x_11, x_12, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_13) == 0) +{ +uint8_t x_14; +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +return x_13; } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___closed__3() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes(" to ", 4); -return x_1; +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +lean_inc(x_15); +lean_dec(x_13); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set(x_17, 1, x_16); +return x_17; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___closed__4() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___closed__3; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +uint8_t x_18; +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; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___closed__5() { +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_formatTable___spec__4___closed__1; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +lean_object* x_13; +x_13 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___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); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +return x_13; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___closed__6() { +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___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_1; -x_1 = lean_mk_string_from_bytes("at ", 3); -return x_1; +lean_object* x_13; +x_13 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +return x_13; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___closed__7() { +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___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_object* x_14) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___closed__6; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +lean_object* x_15; +x_15 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +return x_15; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___closed__8() { +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes(":\n", 2); -return x_1; +lean_object* x_12; +x_12 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__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_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +return x_12; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___closed__9() { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___closed__8; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +lean_object* x_9; +x_9 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___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; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___closed__10() { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___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) { _start: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___lambda__1___boxed), 8, 0); -return x_1; +size_t x_12; size_t x_13; lean_object* x_14; +x_12 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_13 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_14 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5(x_1, x_2, x_3, x_12, x_13, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_14; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___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_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_explainMutualFailure___boxed(lean_object* x_1, lean_object* x_2, 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_12; -x_12 = lean_usize_dec_lt(x_5, x_4); -if (x_12 == 0) -{ -lean_object* x_13; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_6); -lean_ctor_set(x_13, 1, x_11); -return x_13; +lean_object* x_9; +x_9 = l_Lean_Elab_WF_GuessLex_explainMutualFailure(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; } -else -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; 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; uint8_t x_30; size_t x_31; lean_object* x_32; uint8_t x_33; lean_object* x_34; -x_14 = lean_array_uget(x_3, x_5); -x_15 = lean_ctor_get(x_14, 1); -lean_inc(x_15); -x_16 = lean_ctor_get(x_15, 1); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 3); -lean_inc(x_17); -x_18 = l_Lean_Elab_WF_GuessLex_RecCallWithContext_posString(x_15, 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_ctor_get(x_18, 1); -lean_inc(x_20); -lean_dec(x_18); -x_21 = lean_array_get_size(x_1); -x_22 = lean_nat_dec_lt(x_16, x_21); -x_23 = lean_nat_dec_lt(x_17, x_21); -lean_dec(x_21); -x_24 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_24, 0, x_19); -x_25 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___closed__7; -x_26 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_24); -x_27 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___closed__9; -x_28 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = lean_array_get_size(x_2); -x_30 = lean_nat_dec_lt(x_16, x_29); -x_31 = 0; -x_32 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___closed__10; -x_33 = lean_nat_dec_eq(x_16, x_17); -if (x_22 == 0) -{ -lean_object* x_121; lean_object* x_122; -x_121 = l_Lean_instInhabitedName; -x_122 = l___private_Init_Util_0__outOfBounds___rarg(x_121); -x_34 = x_122; -goto block_120; } -else +static lean_object* _init_l_Lean_Elab_WF_GuessLex_explainFailure___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("The arguments relate at each recursive call as follows:\n", 56); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_WF_GuessLex_explainFailure___closed__2() { +_start: { -lean_object* x_123; -x_123 = lean_array_fget(x_1, x_16); -x_34 = x_123; -goto block_120; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_WF_GuessLex_explainFailure___closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } -block_120: +} +static lean_object* _init_l_Lean_Elab_WF_GuessLex_explainFailure___closed__3() { +_start: { -uint8_t x_35; lean_object* x_36; 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_35 = 1; -x_36 = l_Lean_Name_toString(x_34, x_35); -x_37 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_37, 0, x_36); -x_38 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___closed__2; -x_39 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_37); -x_40 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___closed__4; -x_41 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_41, 0, x_39); -lean_ctor_set(x_41, 1, x_40); -if (x_23 == 0) +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("(<, ≤, =: relation proved, \? all proofs failed, _: no proof attempted)\n", 73); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_WF_GuessLex_explainFailure___closed__4() { +_start: { -lean_object* x_117; lean_object* x_118; -x_117 = l_Lean_instInhabitedName; -x_118 = l___private_Init_Util_0__outOfBounds___rarg(x_117); -x_42 = x_118; -goto block_116; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_WF_GuessLex_explainFailure___closed__3; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } -else +} +static lean_object* _init_l_Lean_Elab_WF_GuessLex_explainFailure___closed__5() { +_start: { -lean_object* x_119; -x_119 = lean_array_fget(x_1, x_17); -x_42 = x_119; -goto block_116; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_WF_GuessLex_explainFailure___closed__2; +x_2 = l_Lean_Elab_WF_GuessLex_explainFailure___closed__4; +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; } -block_116: +} +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_explainFailure(lean_object* x_1, lean_object* x_2, 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_43; lean_object* x_44; 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_43 = l_Lean_Name_toString(x_42, x_35); -x_44 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_44, 0, x_43); -x_45 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_45, 0, x_41); -lean_ctor_set(x_45, 1, x_44); -x_46 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___closed__5; -x_47 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_47, 0, x_45); -lean_ctor_set(x_47, 1, x_46); -x_48 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_48, 0, x_6); -lean_ctor_set(x_48, 1, x_47); -x_49 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_49, 0, x_48); -lean_ctor_set(x_49, 1, x_28); -if (x_30 == 0) +lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_9 = lean_array_get_size(x_1); +x_10 = lean_unsigned_to_nat(1u); +x_11 = lean_nat_dec_eq(x_9, x_10); +lean_dec(x_9); +if (x_11 == 0) { -lean_object* x_113; lean_object* x_114; -lean_dec(x_16); -x_113 = l_Lean_Elab_WF_GuessLex_naryVarNames___closed__1; -x_114 = l___private_Init_Util_0__outOfBounds___rarg(x_113); -x_50 = x_114; -goto block_112; +lean_object* x_12; +x_12 = l_Lean_Elab_WF_GuessLex_explainMutualFailure(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_3); +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; lean_object* x_16; +x_14 = lean_ctor_get(x_12, 0); +x_15 = l_Lean_Elab_WF_GuessLex_explainFailure___closed__5; +x_16 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_14); +lean_ctor_set(x_12, 0, x_16); +return x_12; } else { -lean_object* x_115; -x_115 = lean_array_fget(x_2, x_16); -lean_dec(x_16); -x_50 = x_115; -goto block_112; +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_17 = lean_ctor_get(x_12, 0); +x_18 = lean_ctor_get(x_12, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_12); +x_19 = l_Lean_Elab_WF_GuessLex_explainFailure___closed__5; +x_20 = lean_alloc_ctor(5, 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_18); +return x_21; } -block_112: -{ -lean_object* x_51; size_t x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_51 = lean_array_get_size(x_50); -x_52 = lean_usize_of_nat(x_51); -x_53 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_explainNonMutualFailure___spec__1(x_52, x_31, x_50); -x_54 = l_Lean_Elab_WF_GuessLex_explainNonMutualFailure___closed__1; -x_55 = l_Array_append___rarg(x_54, x_53); -x_56 = l_Lean_Elab_WF_GuessLex_GuessLexRel_toNatRel___closed__20; -x_57 = lean_array_push(x_56, x_55); -if (x_33 == 0) +} +else { -uint8_t x_84; -x_84 = lean_nat_dec_lt(x_17, x_29); -lean_dec(x_29); -if (x_84 == 0) +uint8_t x_22; +x_22 = !lean_is_exclusive(x_12); +if (x_22 == 0) { -lean_object* x_85; lean_object* x_86; -lean_dec(x_17); -x_85 = l_Lean_Elab_WF_GuessLex_naryVarNames___closed__1; -x_86 = l___private_Init_Util_0__outOfBounds___rarg(x_85); -x_58 = x_86; -goto block_83; +return x_12; } else { -lean_object* x_87; -x_87 = lean_array_fget(x_2, x_17); -lean_dec(x_17); -x_58 = x_87; -goto block_83; +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_12, 0); +x_24 = lean_ctor_get(x_12, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_12); +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_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; -lean_dec(x_29); -lean_dec(x_17); -x_88 = lean_unsigned_to_nat(0u); -x_89 = lean_unsigned_to_nat(1u); -lean_inc(x_51); -x_90 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__4(x_14, x_51, x_88, x_51, x_89, x_54, x_7, x_8, x_9, x_10, x_20); -lean_dec(x_51); -lean_dec(x_14); -x_91 = lean_ctor_get(x_90, 0); -lean_inc(x_91); -x_92 = lean_ctor_get(x_90, 1); -lean_inc(x_92); -lean_dec(x_90); -x_93 = lean_array_push(x_57, x_91); -x_94 = lean_box(0); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_95 = lean_apply_8(x_32, x_49, x_93, x_94, x_7, x_8, x_9, x_10, x_92); -if (lean_obj_tag(x_95) == 0) +lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_26 = lean_array_get_size(x_2); +x_27 = lean_unsigned_to_nat(0u); +x_28 = lean_nat_dec_lt(x_27, x_26); +lean_dec(x_26); +if (x_28 == 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_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_29 = l_Lean_Elab_WF_GuessLex_naryVarNames___closed__1; +x_30 = l___private_Init_Util_0__outOfBounds___rarg(x_29); +x_31 = l_Lean_Elab_WF_GuessLex_explainNonMutualFailure(x_30, 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); +x_32 = !lean_is_exclusive(x_31); +if (x_32 == 0) { -uint8_t x_97; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_31, 0); +x_34 = l_Lean_Elab_WF_GuessLex_explainFailure___closed__5; +x_35 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_33); +lean_ctor_set(x_31, 0, x_35); +return x_31; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_36 = lean_ctor_get(x_31, 0); +x_37 = lean_ctor_get(x_31, 1); +lean_inc(x_37); +lean_inc(x_36); +lean_dec(x_31); +x_38 = l_Lean_Elab_WF_GuessLex_explainFailure___closed__5; +x_39 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_36); +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_37); +return x_40; +} +} +else +{ +lean_object* x_41; lean_object* x_42; uint8_t x_43; +x_41 = lean_array_fget(x_2, x_27); +x_42 = l_Lean_Elab_WF_GuessLex_explainNonMutualFailure(x_41, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_7); -x_97 = !lean_is_exclusive(x_95); -if (x_97 == 0) +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_43 = !lean_is_exclusive(x_42); +if (x_43 == 0) { -lean_object* x_98; lean_object* x_99; -x_98 = lean_ctor_get(x_95, 0); -lean_dec(x_98); -x_99 = lean_ctor_get(x_96, 0); -lean_inc(x_99); -lean_dec(x_96); -lean_ctor_set(x_95, 0, x_99); -return x_95; +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_42, 0); +x_45 = l_Lean_Elab_WF_GuessLex_explainFailure___closed__5; +x_46 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_44); +lean_ctor_set(x_42, 0, x_46); +return x_42; } 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_ctor_get(x_96, 0); -lean_inc(x_101); -lean_dec(x_96); -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_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_47 = lean_ctor_get(x_42, 0); +x_48 = lean_ctor_get(x_42, 1); +lean_inc(x_48); +lean_inc(x_47); +lean_dec(x_42); +x_49 = l_Lean_Elab_WF_GuessLex_explainFailure___closed__5; +x_50 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_50, 0, x_49); +lean_ctor_set(x_50, 1, x_47); +x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_48); +return x_51; } } -else +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_explainFailure___boxed(lean_object* x_1, lean_object* x_2, 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_103; lean_object* x_104; size_t x_105; size_t x_106; -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 = 1; -x_106 = lean_usize_add(x_5, x_105); -x_5 = x_106; -x_6 = x_104; -x_11 = x_103; -goto _start; +lean_object* x_9; +x_9 = l_Lean_Elab_WF_GuessLex_explainFailure(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_2); +lean_dec(x_1); +return x_9; } } -else +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_reportWF___spec__1(size_t x_1, size_t x_2, lean_object* x_3) { +_start: { -uint8_t x_108; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -x_108 = !lean_is_exclusive(x_95); -if (x_108 == 0) +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) { -return x_95; +return x_3; } else { -lean_object* x_109; lean_object* x_110; lean_object* x_111; -x_109 = lean_ctor_get(x_95, 0); -x_110 = lean_ctor_get(x_95, 1); -lean_inc(x_110); -lean_inc(x_109); -lean_dec(x_95); -x_111 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_111, 0, x_109); -lean_ctor_set(x_111, 1, x_110); -return x_111; +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; size_t x_10; size_t x_11; lean_object* x_12; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = lean_ctor_get(x_5, 6); +lean_inc(x_8); +lean_dec(x_5); +x_9 = lean_ctor_get(x_8, 4); +lean_inc(x_9); +lean_dec(x_8); +x_10 = 1; +x_11 = lean_usize_add(x_2, x_10); +x_12 = lean_array_uset(x_7, x_2, x_9); +x_2 = x_11; +x_3 = x_12; +goto _start; } } } -block_83: +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___lambda__1___closed__1() { +_start: { -lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_59 = lean_array_get_size(x_58); -x_60 = lean_unsigned_to_nat(0u); -x_61 = lean_unsigned_to_nat(1u); -lean_inc(x_59); -x_62 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__3(x_14, x_51, x_58, x_59, x_59, x_60, x_59, x_61, x_57, x_7, x_8, x_9, x_10, x_20); -lean_dec(x_59); -lean_dec(x_58); -lean_dec(x_14); -x_63 = lean_ctor_get(x_62, 0); -lean_inc(x_63); -x_64 = lean_ctor_get(x_62, 1); -lean_inc(x_64); -lean_dec(x_62); -x_65 = lean_box(0); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_66 = lean_apply_8(x_32, x_49, x_63, x_65, x_7, x_8, x_9, x_10, x_64); -if (lean_obj_tag(x_66) == 0) +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Termination", 11); +return x_1; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___lambda__1___closed__2() { +_start: { -lean_object* x_67; -x_67 = lean_ctor_get(x_66, 0); -lean_inc(x_67); -if (lean_obj_tag(x_67) == 0) +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("terminationBy", 13); +return x_1; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___lambda__1___closed__3() { +_start: { -uint8_t x_68; +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_WF_GuessLex_initFn____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_9____closed__6; +x_2 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__1; +x_3 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___lambda__1___closed__1; +x_4 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___lambda__1___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_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___lambda__1___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Try this: ", 10); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___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) { +_start: +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_ctor_get(x_1, 6); +lean_inc(x_10); +lean_dec(x_1); +x_11 = lean_ctor_get(x_10, 1); +lean_inc(x_11); lean_dec(x_10); -lean_dec(x_9); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_12; lean_object* x_13; lean_dec(x_8); lean_dec(x_7); -x_68 = !lean_is_exclusive(x_66); -if (x_68 == 0) -{ -lean_object* x_69; lean_object* x_70; -x_69 = lean_ctor_get(x_66, 0); -lean_dec(x_69); -x_70 = lean_ctor_get(x_67, 0); -lean_inc(x_70); -lean_dec(x_67); -lean_ctor_set(x_66, 0, x_70); -return x_66; +lean_dec(x_3); +x_12 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_12, 0, x_2); +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; } else { -lean_object* x_71; lean_object* x_72; lean_object* x_73; -x_71 = lean_ctor_get(x_66, 1); -lean_inc(x_71); -lean_dec(x_66); -x_72 = lean_ctor_get(x_67, 0); -lean_inc(x_72); -lean_dec(x_67); -x_73 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_73, 0, x_72); -lean_ctor_set(x_73, 1, x_71); -return x_73; -} -} -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; +x_14 = lean_ctor_get(x_11, 0); +lean_inc(x_14); +lean_dec(x_11); +lean_inc(x_7); +x_15 = l_Lean_Elab_WF_TerminationBy_unexpand(x_3, 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_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___lambda__1___closed__3; +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_16); +x_20 = lean_box(0); +x_21 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +lean_ctor_set(x_21, 2, x_20); +lean_ctor_set(x_21, 3, x_20); +lean_ctor_set(x_21, 4, x_20); +lean_ctor_set(x_21, 5, x_20); +x_22 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___lambda__1___closed__4; +x_23 = l_Lean_Meta_Tactic_TryThis_addSuggestion(x_14, x_21, x_20, x_22, x_20, x_5, x_6, x_7, x_8, x_17); +if (lean_obj_tag(x_23) == 0) { -lean_object* x_74; lean_object* x_75; size_t x_76; size_t x_77; -x_74 = lean_ctor_get(x_66, 1); -lean_inc(x_74); -lean_dec(x_66); -x_75 = lean_ctor_get(x_67, 0); -lean_inc(x_75); -lean_dec(x_67); -x_76 = 1; -x_77 = lean_usize_add(x_5, x_76); -x_5 = x_77; -x_6 = x_75; -x_11 = x_74; -goto _start; +uint8_t x_24; +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; +x_25 = lean_ctor_get(x_23, 0); +lean_dec(x_25); +x_26 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_26, 0, x_2); +lean_ctor_set(x_23, 0, x_26); +return x_23; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_23, 1); +lean_inc(x_27); +lean_dec(x_23); +x_28 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_28, 0, x_2); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_27); +return x_29; } } else { -uint8_t x_79; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -x_79 = !lean_is_exclusive(x_66); -if (x_79 == 0) +uint8_t x_30; +lean_dec(x_2); +x_30 = !lean_is_exclusive(x_23); +if (x_30 == 0) { -return x_66; +return x_23; } else { -lean_object* x_80; lean_object* x_81; lean_object* x_82; -x_80 = lean_ctor_get(x_66, 0); -x_81 = lean_ctor_get(x_66, 1); -lean_inc(x_81); -lean_inc(x_80); -lean_dec(x_66); -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_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; } } } } } +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Elab_WF_GuessLex_showInferredTerminationBy; +return x_1; } } +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Inferred termination argument:\n", 31); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_explainMutualFailure(lean_object* x_1, lean_object* x_2, 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_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___closed__3() { _start: { -lean_object* x_9; size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; -x_9 = lean_array_get_size(x_3); -x_10 = lean_usize_of_nat(x_9); -lean_dec(x_9); -x_11 = 0; -x_12 = lean_box(0); -x_13 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5(x_1, x_2, x_3, x_10, x_11, x_12, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_13) == 0) +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___closed__2; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___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) { +_start: { -uint8_t x_14; -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) +uint8_t x_10; +x_10 = lean_usize_dec_lt(x_3, x_2); +if (x_10 == 0) { -return x_13; +lean_object* x_11; +lean_dec(x_8); +lean_dec(x_7); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_4); +lean_ctor_set(x_11, 1, x_9); +return x_11; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_13, 0); -x_16 = lean_ctor_get(x_13, 1); -lean_inc(x_16); +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_12 = lean_array_uget(x_1, x_3); +x_13 = lean_ctor_get(x_4, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_4, 1); +lean_inc(x_14); +x_15 = lean_ctor_get(x_4, 2); lean_inc(x_15); +x_16 = lean_nat_dec_lt(x_14, x_15); +if (x_16 == 0) +{ +lean_object* x_17; +lean_dec(x_15); +lean_dec(x_14); lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_8); +lean_dec(x_7); 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_17, 0, x_4); +lean_ctor_set(x_17, 1, x_9); return x_17; } -} else { uint8_t x_18; -x_18 = !lean_is_exclusive(x_13); +x_18 = !lean_is_exclusive(x_4); if (x_18 == 0) { -return x_13; -} -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; uint8_t x_27; +x_19 = lean_ctor_get(x_4, 2); +lean_dec(x_19); +x_20 = lean_ctor_get(x_4, 1); +lean_dec(x_20); +x_21 = lean_ctor_get(x_4, 0); +lean_dec(x_21); +x_22 = lean_array_fget(x_13, x_14); +x_23 = lean_unsigned_to_nat(1u); +x_24 = lean_nat_add(x_14, x_23); +lean_dec(x_14); +lean_ctor_set(x_4, 1, x_24); +x_25 = lean_ctor_get(x_7, 2); +lean_inc(x_25); +x_26 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___closed__1; +x_27 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_25, x_26); +lean_dec(x_25); +if (x_27 == 0) { -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; -} -} -} -} -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: +lean_object* x_28; lean_object* x_29; +x_28 = lean_box(0); +lean_inc(x_8); +lean_inc(x_7); +x_29 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___lambda__1(x_12, x_4, x_22, x_28, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_29) == 0) { -lean_object* x_13; -x_13 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___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); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); +lean_object* x_30; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +if (lean_obj_tag(x_30) == 0) +{ +uint8_t x_31; lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_2); -lean_dec(x_1); -return x_13; +lean_dec(x_7); +x_31 = !lean_is_exclusive(x_29); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; +x_32 = lean_ctor_get(x_29, 0); +lean_dec(x_32); +x_33 = lean_ctor_get(x_30, 0); +lean_inc(x_33); +lean_dec(x_30); +lean_ctor_set(x_29, 0, x_33); +return x_29; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_29, 1); +lean_inc(x_34); +lean_dec(x_29); +x_35 = lean_ctor_get(x_30, 0); +lean_inc(x_35); +lean_dec(x_30); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_34); +return x_36; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___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 { -lean_object* x_13; -x_13 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_2); -lean_dec(x_1); -return x_13; +lean_object* x_37; lean_object* x_38; size_t x_39; size_t x_40; +x_37 = lean_ctor_get(x_29, 1); +lean_inc(x_37); +lean_dec(x_29); +x_38 = lean_ctor_get(x_30, 0); +lean_inc(x_38); +lean_dec(x_30); +x_39 = 1; +x_40 = lean_usize_add(x_3, x_39); +x_3 = x_40; +x_4 = x_38; +x_9 = x_37; +goto _start; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___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_object* x_14) { -_start: +else { -lean_object* x_15; -x_15 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); +uint8_t x_42; lean_dec(x_8); lean_dec(x_7); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -return x_15; +x_42 = !lean_is_exclusive(x_29); +if (x_42 == 0) +{ +return x_29; } +else +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_43 = lean_ctor_get(x_29, 0); +x_44 = lean_ctor_get(x_29, 1); +lean_inc(x_44); +lean_inc(x_43); +lean_dec(x_29); +x_45 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +return x_45; } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; -x_12 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__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_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -return x_12; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: +else { -lean_object* x_9; -x_9 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +lean_inc(x_7); +lean_inc(x_22); +x_46 = l_Lean_Elab_WF_TerminationBy_unexpand(x_22, x_5, x_6, x_7, x_8, x_9); +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_ctor_get(x_12, 0); +lean_inc(x_49); +x_50 = l_Lean_MessageData_ofSyntax(x_47); +x_51 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___closed__3; +x_52 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_50); +x_53 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__21___rarg___lambda__2___closed__5; +x_54 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +x_55 = 0; +x_56 = l_Lean_logAt___at_Lean_Meta_computeSynthOrder___spec__7(x_49, x_54, x_55, x_5, x_6, x_7, x_8, x_48); +lean_dec(x_49); +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_56, 1); +lean_inc(x_58); +lean_dec(x_56); +lean_inc(x_8); +lean_inc(x_7); +x_59 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___lambda__1(x_12, x_4, x_22, x_57, x_5, x_6, x_7, x_8, x_58); +lean_dec(x_57); +if (lean_obj_tag(x_59) == 0) +{ +lean_object* x_60; +x_60 = lean_ctor_get(x_59, 0); +lean_inc(x_60); +if (lean_obj_tag(x_60) == 0) +{ +uint8_t 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); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___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) { -_start: +x_61 = !lean_is_exclusive(x_59); +if (x_61 == 0) { -size_t x_12; size_t x_13; lean_object* x_14; -x_12 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_13 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_14 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__5(x_1, x_2, x_3, x_12, x_13, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_14; -} +lean_object* x_62; lean_object* x_63; +x_62 = lean_ctor_get(x_59, 0); +lean_dec(x_62); +x_63 = lean_ctor_get(x_60, 0); +lean_inc(x_63); +lean_dec(x_60); +lean_ctor_set(x_59, 0, x_63); +return x_59; } -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_explainMutualFailure___boxed(lean_object* x_1, lean_object* x_2, 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_WF_GuessLex_explainMutualFailure(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; +lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_64 = lean_ctor_get(x_59, 1); +lean_inc(x_64); +lean_dec(x_59); +x_65 = lean_ctor_get(x_60, 0); +lean_inc(x_65); +lean_dec(x_60); +x_66 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_66, 0, x_65); +lean_ctor_set(x_66, 1, x_64); +return x_66; } } -static lean_object* _init_l_Lean_Elab_WF_GuessLex_explainFailure___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("The arguments relate at each recursive call as follows:\n", 56); -return x_1; +lean_object* x_67; lean_object* x_68; size_t x_69; size_t x_70; +x_67 = lean_ctor_get(x_59, 1); +lean_inc(x_67); +lean_dec(x_59); +x_68 = lean_ctor_get(x_60, 0); +lean_inc(x_68); +lean_dec(x_60); +x_69 = 1; +x_70 = lean_usize_add(x_3, x_69); +x_3 = x_70; +x_4 = x_68; +x_9 = x_67; +goto _start; } } -static lean_object* _init_l_Lean_Elab_WF_GuessLex_explainFailure___closed__2() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_WF_GuessLex_explainFailure___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_Elab_WF_GuessLex_explainFailure___closed__3() { -_start: +uint8_t x_72; +lean_dec(x_8); +lean_dec(x_7); +x_72 = !lean_is_exclusive(x_59); +if (x_72 == 0) { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("(<, ≤, =: relation proved, \? all proofs failed, _: no proof attempted)\n", 73); -return x_1; -} +return x_59; } -static lean_object* _init_l_Lean_Elab_WF_GuessLex_explainFailure___closed__4() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_WF_GuessLex_explainFailure___closed__3; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_73 = lean_ctor_get(x_59, 0); +x_74 = lean_ctor_get(x_59, 1); +lean_inc(x_74); +lean_inc(x_73); +lean_dec(x_59); +x_75 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_75, 0, x_73); +lean_ctor_set(x_75, 1, x_74); +return x_75; } } -static lean_object* _init_l_Lean_Elab_WF_GuessLex_explainFailure___closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_WF_GuessLex_explainFailure___closed__2; -x_2 = l_Lean_Elab_WF_GuessLex_explainFailure___closed__4; -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; } } -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_explainFailure(lean_object* x_1, lean_object* x_2, 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; -x_9 = lean_array_get_size(x_1); -x_10 = lean_unsigned_to_nat(1u); -x_11 = lean_nat_dec_eq(x_9, x_10); -lean_dec(x_9); -if (x_11 == 0) +else { -lean_object* x_12; -x_12 = l_Lean_Elab_WF_GuessLex_explainMutualFailure(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_3); -if (lean_obj_tag(x_12) == 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; uint8_t x_82; +lean_dec(x_4); +x_76 = lean_array_fget(x_13, x_14); +x_77 = lean_unsigned_to_nat(1u); +x_78 = lean_nat_add(x_14, x_77); +lean_dec(x_14); +x_79 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_79, 0, x_13); +lean_ctor_set(x_79, 1, x_78); +lean_ctor_set(x_79, 2, x_15); +x_80 = lean_ctor_get(x_7, 2); +lean_inc(x_80); +x_81 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___closed__1; +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) { -uint8_t x_13; -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) +lean_object* x_83; lean_object* x_84; +x_83 = lean_box(0); +lean_inc(x_8); +lean_inc(x_7); +x_84 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___lambda__1(x_12, x_79, x_76, x_83, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_84) == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_12, 0); -x_15 = l_Lean_Elab_WF_GuessLex_explainFailure___closed__5; -x_16 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_14); -lean_ctor_set(x_12, 0, x_16); -return x_12; -} -else +lean_object* x_85; +x_85 = lean_ctor_get(x_84, 0); +lean_inc(x_85); +if (lean_obj_tag(x_85) == 0) { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_17 = lean_ctor_get(x_12, 0); -x_18 = lean_ctor_get(x_12, 1); -lean_inc(x_18); -lean_inc(x_17); -lean_dec(x_12); -x_19 = l_Lean_Elab_WF_GuessLex_explainFailure___closed__5; -x_20 = lean_alloc_ctor(5, 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_18); -return x_21; +lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; +lean_dec(x_8); +lean_dec(x_7); +x_86 = lean_ctor_get(x_84, 1); +lean_inc(x_86); +if (lean_is_exclusive(x_84)) { + lean_ctor_release(x_84, 0); + lean_ctor_release(x_84, 1); + x_87 = x_84; +} else { + lean_dec_ref(x_84); + x_87 = lean_box(0); } +x_88 = lean_ctor_get(x_85, 0); +lean_inc(x_88); +lean_dec(x_85); +if (lean_is_scalar(x_87)) { + x_89 = lean_alloc_ctor(0, 2, 0); +} else { + x_89 = x_87; +} +lean_ctor_set(x_89, 0, x_88); +lean_ctor_set(x_89, 1, x_86); +return x_89; } else { -uint8_t x_22; -x_22 = !lean_is_exclusive(x_12); -if (x_22 == 0) -{ -return x_12; +lean_object* x_90; lean_object* x_91; size_t x_92; size_t x_93; +x_90 = lean_ctor_get(x_84, 1); +lean_inc(x_90); +lean_dec(x_84); +x_91 = lean_ctor_get(x_85, 0); +lean_inc(x_91); +lean_dec(x_85); +x_92 = 1; +x_93 = lean_usize_add(x_3, x_92); +x_3 = x_93; +x_4 = x_91; +x_9 = x_90; +goto _start; +} } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_12, 0); -x_24 = lean_ctor_get(x_12, 1); -lean_inc(x_24); -lean_inc(x_23); -lean_dec(x_12); -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_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; +lean_dec(x_8); +lean_dec(x_7); +x_95 = lean_ctor_get(x_84, 0); +lean_inc(x_95); +x_96 = lean_ctor_get(x_84, 1); +lean_inc(x_96); +if (lean_is_exclusive(x_84)) { + lean_ctor_release(x_84, 0); + lean_ctor_release(x_84, 1); + x_97 = x_84; +} else { + lean_dec_ref(x_84); + x_97 = lean_box(0); +} +if (lean_is_scalar(x_97)) { + x_98 = lean_alloc_ctor(1, 2, 0); +} else { + x_98 = x_97; } +lean_ctor_set(x_98, 0, x_95); +lean_ctor_set(x_98, 1, x_96); +return x_98; } } else { -lean_object* x_26; lean_object* x_27; uint8_t x_28; -x_26 = lean_array_get_size(x_2); -x_27 = lean_unsigned_to_nat(0u); -x_28 = lean_nat_dec_lt(x_27, x_26); -lean_dec(x_26); -if (x_28 == 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_object* x_105; lean_object* x_106; lean_object* x_107; uint8_t x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +lean_inc(x_7); +lean_inc(x_76); +x_99 = l_Lean_Elab_WF_TerminationBy_unexpand(x_76, x_5, x_6, x_7, x_8, x_9); +x_100 = lean_ctor_get(x_99, 0); +lean_inc(x_100); +x_101 = lean_ctor_get(x_99, 1); +lean_inc(x_101); +lean_dec(x_99); +x_102 = lean_ctor_get(x_12, 0); +lean_inc(x_102); +x_103 = l_Lean_MessageData_ofSyntax(x_100); +x_104 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___closed__3; +x_105 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_105, 0, x_104); +lean_ctor_set(x_105, 1, x_103); +x_106 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__21___rarg___lambda__2___closed__5; +x_107 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_107, 0, x_105); +lean_ctor_set(x_107, 1, x_106); +x_108 = 0; +x_109 = l_Lean_logAt___at_Lean_Meta_computeSynthOrder___spec__7(x_102, x_107, x_108, x_5, x_6, x_7, x_8, x_101); +lean_dec(x_102); +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_8); +lean_inc(x_7); +x_112 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___lambda__1(x_12, x_79, x_76, x_110, x_5, x_6, x_7, x_8, x_111); +lean_dec(x_110); +if (lean_obj_tag(x_112) == 0) { -lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; -x_29 = l_Lean_Elab_WF_GuessLex_naryVarNames___closed__1; -x_30 = l___private_Init_Util_0__outOfBounds___rarg(x_29); -x_31 = l_Lean_Elab_WF_GuessLex_explainNonMutualFailure(x_30, 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); -x_32 = !lean_is_exclusive(x_31); -if (x_32 == 0) +lean_object* x_113; +x_113 = lean_ctor_get(x_112, 0); +lean_inc(x_113); +if (lean_obj_tag(x_113) == 0) { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_31, 0); -x_34 = l_Lean_Elab_WF_GuessLex_explainFailure___closed__5; -x_35 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_33); -lean_ctor_set(x_31, 0, x_35); -return x_31; +lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; +lean_dec(x_8); +lean_dec(x_7); +x_114 = lean_ctor_get(x_112, 1); +lean_inc(x_114); +if (lean_is_exclusive(x_112)) { + lean_ctor_release(x_112, 0); + lean_ctor_release(x_112, 1); + x_115 = x_112; +} else { + lean_dec_ref(x_112); + x_115 = lean_box(0); +} +x_116 = lean_ctor_get(x_113, 0); +lean_inc(x_116); +lean_dec(x_113); +if (lean_is_scalar(x_115)) { + x_117 = lean_alloc_ctor(0, 2, 0); +} else { + x_117 = x_115; +} +lean_ctor_set(x_117, 0, x_116); +lean_ctor_set(x_117, 1, x_114); +return x_117; } else { -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_36 = lean_ctor_get(x_31, 0); -x_37 = lean_ctor_get(x_31, 1); -lean_inc(x_37); -lean_inc(x_36); -lean_dec(x_31); -x_38 = l_Lean_Elab_WF_GuessLex_explainFailure___closed__5; -x_39 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_36); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_39); -lean_ctor_set(x_40, 1, x_37); -return x_40; +lean_object* x_118; lean_object* x_119; size_t x_120; size_t x_121; +x_118 = lean_ctor_get(x_112, 1); +lean_inc(x_118); +lean_dec(x_112); +x_119 = lean_ctor_get(x_113, 0); +lean_inc(x_119); +lean_dec(x_113); +x_120 = 1; +x_121 = lean_usize_add(x_3, x_120); +x_3 = x_121; +x_4 = x_119; +x_9 = x_118; +goto _start; } } else { -lean_object* x_41; lean_object* x_42; uint8_t x_43; -x_41 = lean_array_fget(x_2, x_27); -x_42 = l_Lean_Elab_WF_GuessLex_explainNonMutualFailure(x_41, x_3, x_4, x_5, x_6, x_7, x_8); +lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; +lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_43 = !lean_is_exclusive(x_42); -if (x_43 == 0) +x_123 = lean_ctor_get(x_112, 0); +lean_inc(x_123); +x_124 = lean_ctor_get(x_112, 1); +lean_inc(x_124); +if (lean_is_exclusive(x_112)) { + lean_ctor_release(x_112, 0); + lean_ctor_release(x_112, 1); + x_125 = x_112; +} else { + lean_dec_ref(x_112); + x_125 = lean_box(0); +} +if (lean_is_scalar(x_125)) { + x_126 = lean_alloc_ctor(1, 2, 0); +} else { + x_126 = x_125; +} +lean_ctor_set(x_126, 0, x_123); +lean_ctor_set(x_126, 1, x_124); +return x_126; +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_reportWF(lean_object* x_1, 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_44; lean_object* x_45; lean_object* x_46; -x_44 = lean_ctor_get(x_42, 0); -x_45 = l_Lean_Elab_WF_GuessLex_explainFailure___closed__5; -x_46 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_46, 0, x_45); -lean_ctor_set(x_46, 1, x_44); -lean_ctor_set(x_42, 0, x_46); -return x_42; +lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_8 = lean_array_get_size(x_1); +x_9 = lean_usize_of_nat(x_8); +lean_dec(x_8); +x_10 = 0; +lean_inc(x_1); +x_11 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_reportWF___spec__1(x_9, x_10, x_1); +x_12 = l_Lean_Elab_WF_GuessLex_trimTermWF(x_11, x_2); +lean_dec(x_11); +x_13 = lean_array_get_size(x_12); +x_14 = lean_unsigned_to_nat(0u); +x_15 = l_Array_toSubarray___rarg(x_12, x_14, x_13); +x_16 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2(x_1, x_9, x_10, x_15, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_1); +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; lean_object* x_19; +x_18 = lean_ctor_get(x_16, 0); +lean_dec(x_18); +x_19 = lean_box(0); +lean_ctor_set(x_16, 0, x_19); +return x_16; } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_47 = lean_ctor_get(x_42, 0); -x_48 = lean_ctor_get(x_42, 1); -lean_inc(x_48); -lean_inc(x_47); -lean_dec(x_42); -x_49 = l_Lean_Elab_WF_GuessLex_explainFailure___closed__5; -x_50 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_50, 0, x_49); -lean_ctor_set(x_50, 1, x_47); -x_51 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_51, 0, x_50); -lean_ctor_set(x_51, 1, x_48); -return x_51; +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_16, 1); +lean_inc(x_20); +lean_dec(x_16); +x_21 = lean_box(0); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_20); +return x_22; } } +else +{ +uint8_t x_23; +x_23 = !lean_is_exclusive(x_16); +if (x_23 == 0) +{ +return x_16; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_16, 0); +x_25 = lean_ctor_get(x_16, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_16); +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_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_explainFailure___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_reportWF___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_9; -x_9 = l_Lean_Elab_WF_GuessLex_explainFailure(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_2); +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); lean_dec(x_1); -return x_9; +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_reportWF___spec__1(x_4, x_5, x_3); +return x_6; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__1(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -uint8_t x_4; -x_4 = lean_usize_dec_lt(x_2, x_1); -if (x_4 == 0) -{ -return x_3; +lean_object* x_10; +x_10 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_10; } -else +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___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_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; size_t x_10; size_t x_11; lean_object* x_12; -x_5 = lean_array_uget(x_3, x_2); -x_6 = lean_unsigned_to_nat(0u); -x_7 = lean_array_uset(x_3, x_2, x_6); -x_8 = lean_ctor_get(x_5, 6); -lean_inc(x_8); +size_t x_10; size_t x_11; lean_object* x_12; +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_Elab_WF_GuessLex_reportWF___spec__2(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); -x_9 = lean_ctor_get(x_8, 4); -lean_inc(x_9); -lean_dec(x_8); -x_10 = 1; -x_11 = lean_usize_add(x_2, x_10); -x_12 = lean_array_uset(x_7, x_2, x_9); -x_2 = x_11; -x_3 = x_12; -goto _start; +lean_dec(x_1); +return x_12; } } +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_reportWF___boxed(lean_object* x_1, 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_WF_GuessLex_reportWF(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_8; +} } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__2(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__1(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { uint8_t x_9; @@ -23587,7 +26624,7 @@ return x_24; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___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_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___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) { _start: { uint8_t x_10; @@ -23645,199 +26682,27 @@ lean_dec(x_5); lean_dec(x_1); x_22 = !lean_is_exclusive(x_15); if (x_22 == 0) -{ -return x_15; -} -else -{ -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_15, 0); -x_24 = lean_ctor_get(x_15, 1); -lean_inc(x_24); -lean_inc(x_23); -lean_dec(x_15); -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_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__4(size_t x_1, size_t x_2, lean_object* x_3) { -_start: -{ -uint8_t x_4; -x_4 = lean_usize_dec_lt(x_2, x_1); -if (x_4 == 0) -{ -return x_3; -} -else -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; -x_5 = lean_array_uget(x_3, x_2); -x_6 = lean_unsigned_to_nat(0u); -x_7 = lean_array_uset(x_3, x_2, x_6); -x_8 = lean_array_get_size(x_5); -lean_dec(x_5); -x_9 = 1; -x_10 = lean_usize_add(x_2, x_9); -x_11 = lean_array_uset(x_7, x_2, x_8); -x_2 = x_10; -x_3 = x_11; -goto _start; -} -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___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) { -_start: -{ -uint8_t x_10; -x_10 = lean_usize_dec_lt(x_3, x_2); -if (x_10 == 0) -{ -lean_object* x_11; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_1); -x_11 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_11, 0, x_4); -lean_ctor_set(x_11, 1, x_9); -return x_11; -} -else -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_12 = lean_array_uget(x_4, x_3); -x_13 = lean_unsigned_to_nat(0u); -x_14 = lean_array_uset(x_4, x_3, x_13); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_1); -x_15 = l_Lean_Elab_WF_GuessLex_getForbiddenByTrivialSizeOf(x_1, x_12, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(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_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_20 = lean_array_uset(x_14, x_3, x_16); -x_3 = x_19; -x_4 = x_20; -x_9 = x_17; -goto _start; -} -else -{ -uint8_t x_22; -lean_dec(x_14); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_1); -x_22 = !lean_is_exclusive(x_15); -if (x_22 == 0) -{ -return x_15; -} -else -{ -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_15, 0); -x_24 = lean_ctor_get(x_15, 1); -lean_inc(x_24); -lean_inc(x_23); -lean_dec(x_15); -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_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__6(size_t x_1, size_t x_2, lean_object* x_3) { -_start: -{ -uint8_t x_4; -x_4 = lean_usize_dec_lt(x_2, x_1); -if (x_4 == 0) -{ -return x_3; -} -else -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; size_t x_10; size_t x_11; lean_object* x_12; -x_5 = lean_array_uget(x_3, x_2); -x_6 = lean_unsigned_to_nat(0u); -x_7 = lean_array_uset(x_3, x_2, x_6); -x_8 = lean_ctor_get(x_5, 6); -lean_inc(x_8); -lean_dec(x_5); -x_9 = lean_ctor_get(x_8, 3); -lean_inc(x_9); -lean_dec(x_8); -x_10 = 1; -x_11 = lean_usize_add(x_2, x_10); -x_12 = lean_array_uset(x_7, x_2, x_9); -x_2 = x_11; -x_3 = x_12; -goto _start; -} -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__7(lean_object* x_1, size_t x_2, size_t x_3, size_t x_4, size_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -uint8_t x_12; -x_12 = lean_usize_dec_lt(x_5, x_4); -if (x_12 == 0) -{ -lean_object* x_13; -lean_dec(x_1); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_6); -lean_ctor_set(x_13, 1, x_11); -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; size_t x_21; size_t x_22; lean_object* x_23; -x_14 = lean_array_uget(x_6, x_5); -x_15 = lean_unsigned_to_nat(0u); -x_16 = lean_array_uset(x_6, x_5, x_15); -lean_inc(x_1); -x_17 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__6(x_2, x_3, x_1); -x_18 = l_Lean_Elab_WF_GuessLex_RecCallCache_mk(x_17, x_14, x_11); -lean_dec(x_17); -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_18, 1); -lean_inc(x_20); -lean_dec(x_18); -x_21 = 1; -x_22 = lean_usize_add(x_5, x_21); -x_23 = lean_array_uset(x_16, x_5, x_19); -x_5 = x_22; -x_6 = x_23; -x_11 = x_20; -goto _start; +{ +return x_15; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_15, 0); +x_24 = lean_ctor_get(x_15, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_15); +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_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__8(size_t x_1, size_t x_2, lean_object* x_3) { +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__3(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -23852,8 +26717,8 @@ lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x x_5 = lean_array_uget(x_3, x_2); x_6 = lean_unsigned_to_nat(0u); x_7 = lean_array_uset(x_3, x_2, x_6); -x_8 = lean_alloc_closure((void*)(l_Lean_Elab_WF_GuessLex_inspectCall), 7, 1); -lean_closure_set(x_8, 0, x_5); +x_8 = lean_array_get_size(x_5); +lean_dec(x_5); x_9 = 1; x_10 = lean_usize_add(x_2, x_9); x_11 = lean_array_uset(x_7, x_2, x_8); @@ -23863,1809 +26728,1492 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__11(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_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__4(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -uint8_t x_11; -x_11 = lean_usize_dec_lt(x_4, x_3); -if (x_11 == 0) +uint8_t x_10; +x_10 = lean_usize_dec_lt(x_3, x_2); +if (x_10 == 0) { -lean_object* x_12; -lean_dec(x_9); +lean_object* x_11; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); +lean_dec(x_5); lean_dec(x_1); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_5); -lean_ctor_set(x_12, 1, x_10); -return x_12; +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_4); +lean_ctor_set(x_11, 1, x_9); +return x_11; } else { -lean_object* x_13; uint8_t x_14; -x_13 = lean_array_uget(x_2, x_4); -x_14 = !lean_is_exclusive(x_5); -if (x_14 == 0) -{ -lean_object* x_15; uint8_t x_16; -x_15 = lean_ctor_get(x_5, 1); -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; -x_17 = lean_ctor_get(x_5, 0); -x_18 = lean_ctor_get(x_15, 0); -x_19 = lean_ctor_get(x_15, 1); -x_20 = lean_array_uget(x_2, x_4); -lean_inc(x_9); +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_12 = lean_array_uget(x_4, x_3); +x_13 = lean_unsigned_to_nat(0u); +x_14 = lean_array_uset(x_4, x_3, x_13); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); +lean_inc(x_5); lean_inc(x_1); -x_21 = lean_apply_6(x_20, x_1, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_21) == 0) -{ -uint8_t x_22; -x_22 = !lean_is_exclusive(x_21); -if (x_22 == 0) -{ -lean_object* x_23; lean_object* x_24; uint8_t x_25; uint8_t x_26; uint8_t x_27; -x_23 = lean_ctor_get(x_21, 0); -x_24 = lean_ctor_get(x_21, 1); -x_25 = 0; -x_26 = lean_unbox(x_23); -x_27 = l_Lean_Elab_WF_GuessLex_instDecidableEqGuessLexRel(x_26, x_25); -if (x_27 == 0) -{ -lean_object* x_28; uint8_t x_29; uint8_t x_30; uint8_t x_31; -x_28 = lean_array_push(x_19, x_13); -x_29 = 2; -x_30 = lean_unbox(x_23); -x_31 = l_Lean_Elab_WF_GuessLex_instDecidableEqGuessLexRel(x_30, x_29); -if (x_31 == 0) -{ -uint8_t x_32; uint8_t x_33; uint8_t x_34; -x_32 = 1; -x_33 = lean_unbox(x_23); -lean_dec(x_23); -x_34 = l_Lean_Elab_WF_GuessLex_instDecidableEqGuessLexRel(x_33, x_32); -if (x_34 == 0) -{ -uint8_t x_35; lean_object* x_36; -lean_dec(x_17); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -lean_ctor_set(x_15, 1, x_28); -x_35 = 0; -x_36 = lean_box(x_35); -lean_ctor_set(x_5, 0, x_36); -lean_ctor_set(x_21, 0, x_5); -return x_21; -} -else -{ -size_t x_37; size_t x_38; -lean_free_object(x_21); -lean_ctor_set(x_15, 1, x_28); -x_37 = 1; -x_38 = lean_usize_add(x_4, x_37); -x_4 = x_38; -x_10 = x_24; -goto _start; -} -} -else -{ -size_t x_40; size_t x_41; -lean_free_object(x_21); -lean_dec(x_23); -lean_ctor_set(x_15, 1, x_28); -x_40 = 1; -x_41 = lean_usize_add(x_4, x_40); -x_4 = x_41; -x_10 = x_24; -goto _start; -} -} -else -{ -uint8_t x_43; lean_object* x_44; size_t x_45; size_t x_46; -lean_free_object(x_21); -lean_dec(x_23); -lean_dec(x_18); -lean_dec(x_13); -x_43 = 1; -x_44 = lean_box(x_43); -lean_ctor_set(x_15, 0, x_44); -x_45 = 1; -x_46 = lean_usize_add(x_4, x_45); -x_4 = x_46; -x_10 = x_24; -goto _start; -} -} -else -{ -lean_object* x_48; lean_object* x_49; uint8_t x_50; uint8_t x_51; uint8_t x_52; -x_48 = lean_ctor_get(x_21, 0); -x_49 = lean_ctor_get(x_21, 1); -lean_inc(x_49); -lean_inc(x_48); -lean_dec(x_21); -x_50 = 0; -x_51 = lean_unbox(x_48); -x_52 = l_Lean_Elab_WF_GuessLex_instDecidableEqGuessLexRel(x_51, x_50); -if (x_52 == 0) -{ -lean_object* x_53; uint8_t x_54; uint8_t x_55; uint8_t x_56; -x_53 = lean_array_push(x_19, x_13); -x_54 = 2; -x_55 = lean_unbox(x_48); -x_56 = l_Lean_Elab_WF_GuessLex_instDecidableEqGuessLexRel(x_55, x_54); -if (x_56 == 0) -{ -uint8_t x_57; uint8_t x_58; uint8_t x_59; -x_57 = 1; -x_58 = lean_unbox(x_48); -lean_dec(x_48); -x_59 = l_Lean_Elab_WF_GuessLex_instDecidableEqGuessLexRel(x_58, x_57); -if (x_59 == 0) -{ -uint8_t x_60; lean_object* x_61; lean_object* x_62; -lean_dec(x_17); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -lean_ctor_set(x_15, 1, x_53); -x_60 = 0; -x_61 = lean_box(x_60); -lean_ctor_set(x_5, 0, x_61); -x_62 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_62, 0, x_5); -lean_ctor_set(x_62, 1, x_49); -return x_62; -} -else -{ -size_t x_63; size_t x_64; -lean_ctor_set(x_15, 1, x_53); -x_63 = 1; -x_64 = lean_usize_add(x_4, x_63); -x_4 = x_64; -x_10 = x_49; -goto _start; -} -} -else -{ -size_t x_66; size_t x_67; -lean_dec(x_48); -lean_ctor_set(x_15, 1, x_53); -x_66 = 1; -x_67 = lean_usize_add(x_4, x_66); -x_4 = x_67; -x_10 = x_49; -goto _start; -} -} -else +x_15 = l_Lean_Elab_WF_GuessLex_getForbiddenByTrivialSizeOf(x_1, x_12, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_15) == 0) { -uint8_t x_69; lean_object* x_70; size_t x_71; size_t x_72; -lean_dec(x_48); -lean_dec(x_18); -lean_dec(x_13); -x_69 = 1; -x_70 = lean_box(x_69); -lean_ctor_set(x_15, 0, x_70); -x_71 = 1; -x_72 = lean_usize_add(x_4, x_71); -x_4 = x_72; -x_10 = x_49; +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_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_20 = lean_array_uset(x_14, x_3, x_16); +x_3 = x_19; +x_4 = x_20; +x_9 = x_17; goto _start; } -} -} else { -uint8_t x_74; -lean_free_object(x_15); -lean_dec(x_19); -lean_dec(x_18); -lean_free_object(x_5); -lean_dec(x_17); -lean_dec(x_13); -lean_dec(x_9); +uint8_t x_22; +lean_dec(x_14); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); +lean_dec(x_5); lean_dec(x_1); -x_74 = !lean_is_exclusive(x_21); -if (x_74 == 0) +x_22 = !lean_is_exclusive(x_15); +if (x_22 == 0) { -return x_21; +return x_15; } else { -lean_object* x_75; lean_object* x_76; lean_object* x_77; -x_75 = lean_ctor_get(x_21, 0); -x_76 = lean_ctor_get(x_21, 1); -lean_inc(x_76); -lean_inc(x_75); -lean_dec(x_21); -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_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_15, 0); +x_24 = lean_ctor_get(x_15, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_15); +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_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__5(size_t x_1, size_t x_2, lean_object* x_3) { +_start: { -lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; -x_78 = lean_ctor_get(x_5, 0); -x_79 = lean_ctor_get(x_15, 0); -x_80 = lean_ctor_get(x_15, 1); -lean_inc(x_80); -lean_inc(x_79); -lean_dec(x_15); -x_81 = lean_array_uget(x_2, x_4); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_1); -x_82 = lean_apply_6(x_81, x_1, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_82) == 0) +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) { -lean_object* x_83; lean_object* x_84; lean_object* x_85; uint8_t x_86; uint8_t x_87; uint8_t x_88; -x_83 = lean_ctor_get(x_82, 0); -lean_inc(x_83); -x_84 = lean_ctor_get(x_82, 1); -lean_inc(x_84); -if (lean_is_exclusive(x_82)) { - lean_ctor_release(x_82, 0); - lean_ctor_release(x_82, 1); - x_85 = x_82; -} else { - lean_dec_ref(x_82); - x_85 = lean_box(0); +return x_3; } -x_86 = 0; -x_87 = lean_unbox(x_83); -x_88 = l_Lean_Elab_WF_GuessLex_instDecidableEqGuessLexRel(x_87, x_86); -if (x_88 == 0) +else { -lean_object* x_89; uint8_t x_90; uint8_t x_91; uint8_t x_92; -x_89 = lean_array_push(x_80, x_13); -x_90 = 2; -x_91 = lean_unbox(x_83); -x_92 = l_Lean_Elab_WF_GuessLex_instDecidableEqGuessLexRel(x_91, x_90); -if (x_92 == 0) +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; size_t x_10; size_t x_11; lean_object* x_12; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = lean_ctor_get(x_5, 6); +lean_inc(x_8); +lean_dec(x_5); +x_9 = lean_ctor_get(x_8, 3); +lean_inc(x_9); +lean_dec(x_8); +x_10 = 1; +x_11 = lean_usize_add(x_2, x_10); +x_12 = lean_array_uset(x_7, x_2, x_9); +x_2 = x_11; +x_3 = x_12; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__6(lean_object* x_1, size_t x_2, size_t x_3, size_t x_4, size_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: { -uint8_t x_93; uint8_t x_94; uint8_t x_95; -x_93 = 1; -x_94 = lean_unbox(x_83); -lean_dec(x_83); -x_95 = l_Lean_Elab_WF_GuessLex_instDecidableEqGuessLexRel(x_94, x_93); -if (x_95 == 0) +uint8_t x_12; +x_12 = lean_usize_dec_lt(x_5, x_4); +if (x_12 == 0) { -lean_object* x_96; uint8_t x_97; lean_object* x_98; lean_object* x_99; -lean_dec(x_78); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); +lean_object* x_13; lean_dec(x_1); -x_96 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_96, 0, x_79); -lean_ctor_set(x_96, 1, x_89); -x_97 = 0; -x_98 = lean_box(x_97); -lean_ctor_set(x_5, 1, x_96); -lean_ctor_set(x_5, 0, x_98); -if (lean_is_scalar(x_85)) { - x_99 = lean_alloc_ctor(0, 2, 0); -} else { - x_99 = x_85; -} -lean_ctor_set(x_99, 0, x_5); -lean_ctor_set(x_99, 1, x_84); -return x_99; +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_6); +lean_ctor_set(x_13, 1, x_11); +return x_13; } else { -lean_object* x_100; size_t x_101; size_t x_102; -lean_dec(x_85); -x_100 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_100, 0, x_79); -lean_ctor_set(x_100, 1, x_89); -lean_ctor_set(x_5, 1, x_100); -x_101 = 1; -x_102 = lean_usize_add(x_4, x_101); -x_4 = x_102; -x_10 = x_84; +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; size_t x_21; size_t x_22; lean_object* x_23; +x_14 = lean_array_uget(x_6, x_5); +x_15 = lean_unsigned_to_nat(0u); +x_16 = lean_array_uset(x_6, x_5, x_15); +lean_inc(x_1); +x_17 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__5(x_2, x_3, x_1); +x_18 = l_Lean_Elab_WF_GuessLex_RecCallCache_mk(x_17, x_14, x_11); +lean_dec(x_17); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = 1; +x_22 = lean_usize_add(x_5, x_21); +x_23 = lean_array_uset(x_16, x_5, x_19); +x_5 = x_22; +x_6 = x_23; +x_11 = x_20; goto _start; } } -else -{ -lean_object* x_104; size_t x_105; size_t x_106; -lean_dec(x_85); -lean_dec(x_83); -x_104 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_104, 0, x_79); -lean_ctor_set(x_104, 1, x_89); -lean_ctor_set(x_5, 1, x_104); -x_105 = 1; -x_106 = lean_usize_add(x_4, x_105); -x_4 = x_106; -x_10 = x_84; -goto _start; } +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__7(size_t x_1, size_t x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) +{ +return x_3; } else { -uint8_t x_108; lean_object* x_109; lean_object* x_110; size_t x_111; size_t x_112; -lean_dec(x_85); -lean_dec(x_83); -lean_dec(x_79); -lean_dec(x_13); -x_108 = 1; -x_109 = lean_box(x_108); -x_110 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_110, 0, x_109); -lean_ctor_set(x_110, 1, x_80); -lean_ctor_set(x_5, 1, x_110); -x_111 = 1; -x_112 = lean_usize_add(x_4, x_111); -x_4 = x_112; -x_10 = x_84; +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = lean_alloc_closure((void*)(l_Lean_Elab_WF_GuessLex_inspectCall), 7, 1); +lean_closure_set(x_8, 0, x_5); +x_9 = 1; +x_10 = lean_usize_add(x_2, x_9); +x_11 = lean_array_uset(x_7, x_2, x_8); +x_2 = x_10; +x_3 = x_11; goto _start; } } -else +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___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: { -lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; -lean_dec(x_80); -lean_dec(x_79); -lean_free_object(x_5); -lean_dec(x_78); -lean_dec(x_13); +uint8_t x_11; +x_11 = lean_usize_dec_lt(x_4, x_3); +if (x_11 == 0) +{ +lean_object* x_12; lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_1); -x_114 = lean_ctor_get(x_82, 0); -lean_inc(x_114); -x_115 = lean_ctor_get(x_82, 1); -lean_inc(x_115); -if (lean_is_exclusive(x_82)) { - lean_ctor_release(x_82, 0); - lean_ctor_release(x_82, 1); - x_116 = x_82; -} else { - lean_dec_ref(x_82); - x_116 = lean_box(0); -} -if (lean_is_scalar(x_116)) { - x_117 = lean_alloc_ctor(1, 2, 0); -} else { - x_117 = x_116; -} -lean_ctor_set(x_117, 0, x_114); -lean_ctor_set(x_117, 1, x_115); -return x_117; -} -} +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_5); +lean_ctor_set(x_12, 1, x_10); +return x_12; } else { -lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; -x_118 = lean_ctor_get(x_5, 1); -x_119 = lean_ctor_get(x_5, 0); -lean_inc(x_118); -lean_inc(x_119); -lean_dec(x_5); -x_120 = lean_ctor_get(x_118, 0); -lean_inc(x_120); -x_121 = lean_ctor_get(x_118, 1); -lean_inc(x_121); -if (lean_is_exclusive(x_118)) { - lean_ctor_release(x_118, 0); - lean_ctor_release(x_118, 1); - x_122 = x_118; -} else { - lean_dec_ref(x_118); - x_122 = lean_box(0); -} -x_123 = lean_array_uget(x_2, x_4); +lean_object* x_13; uint8_t x_14; +x_13 = lean_array_uget(x_2, x_4); +x_14 = !lean_is_exclusive(x_5); +if (x_14 == 0) +{ +lean_object* x_15; uint8_t x_16; +x_15 = lean_ctor_get(x_5, 1); +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; +x_17 = lean_ctor_get(x_5, 0); +x_18 = lean_ctor_get(x_15, 0); +x_19 = lean_ctor_get(x_15, 1); +x_20 = lean_array_uget(x_2, x_4); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_1); -x_124 = lean_apply_6(x_123, x_1, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_124) == 0) +x_21 = lean_apply_6(x_20, x_1, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_21) == 0) { -lean_object* x_125; lean_object* x_126; lean_object* x_127; uint8_t x_128; uint8_t x_129; uint8_t x_130; -x_125 = lean_ctor_get(x_124, 0); -lean_inc(x_125); -x_126 = lean_ctor_get(x_124, 1); -lean_inc(x_126); -if (lean_is_exclusive(x_124)) { - lean_ctor_release(x_124, 0); - lean_ctor_release(x_124, 1); - x_127 = x_124; -} else { - lean_dec_ref(x_124); - x_127 = lean_box(0); -} -x_128 = 0; -x_129 = lean_unbox(x_125); -x_130 = l_Lean_Elab_WF_GuessLex_instDecidableEqGuessLexRel(x_129, x_128); -if (x_130 == 0) +uint8_t x_22; +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) { -lean_object* x_131; uint8_t x_132; uint8_t x_133; uint8_t x_134; -x_131 = lean_array_push(x_121, x_13); -x_132 = 2; -x_133 = lean_unbox(x_125); -x_134 = l_Lean_Elab_WF_GuessLex_instDecidableEqGuessLexRel(x_133, x_132); -if (x_134 == 0) +lean_object* x_23; lean_object* x_24; uint8_t x_25; uint8_t x_26; uint8_t x_27; +x_23 = lean_ctor_get(x_21, 0); +x_24 = lean_ctor_get(x_21, 1); +x_25 = 0; +x_26 = lean_unbox(x_23); +x_27 = l_Lean_Elab_WF_GuessLex_instDecidableEqGuessLexRel(x_26, x_25); +if (x_27 == 0) { -uint8_t x_135; uint8_t x_136; uint8_t x_137; -x_135 = 1; -x_136 = lean_unbox(x_125); -lean_dec(x_125); -x_137 = l_Lean_Elab_WF_GuessLex_instDecidableEqGuessLexRel(x_136, x_135); -if (x_137 == 0) +lean_object* x_28; uint8_t x_29; uint8_t x_30; uint8_t x_31; +x_28 = lean_array_push(x_19, x_13); +x_29 = 2; +x_30 = lean_unbox(x_23); +x_31 = l_Lean_Elab_WF_GuessLex_instDecidableEqGuessLexRel(x_30, x_29); +if (x_31 == 0) { -lean_object* x_138; uint8_t x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; -lean_dec(x_119); +uint8_t x_32; uint8_t x_33; uint8_t x_34; +x_32 = 1; +x_33 = lean_unbox(x_23); +lean_dec(x_23); +x_34 = l_Lean_Elab_WF_GuessLex_instDecidableEqGuessLexRel(x_33, x_32); +if (x_34 == 0) +{ +uint8_t x_35; lean_object* x_36; +lean_dec(x_17); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_1); -if (lean_is_scalar(x_122)) { - x_138 = lean_alloc_ctor(0, 2, 0); -} else { - x_138 = x_122; -} -lean_ctor_set(x_138, 0, x_120); -lean_ctor_set(x_138, 1, x_131); -x_139 = 0; -x_140 = lean_box(x_139); -x_141 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_141, 0, x_140); -lean_ctor_set(x_141, 1, x_138); -if (lean_is_scalar(x_127)) { - x_142 = lean_alloc_ctor(0, 2, 0); -} else { - x_142 = x_127; -} -lean_ctor_set(x_142, 0, x_141); -lean_ctor_set(x_142, 1, x_126); -return x_142; +lean_ctor_set(x_15, 1, x_28); +x_35 = 0; +x_36 = lean_box(x_35); +lean_ctor_set(x_5, 0, x_36); +lean_ctor_set(x_21, 0, x_5); +return x_21; } else { -lean_object* x_143; lean_object* x_144; size_t x_145; size_t x_146; -lean_dec(x_127); -if (lean_is_scalar(x_122)) { - x_143 = lean_alloc_ctor(0, 2, 0); -} else { - x_143 = x_122; -} -lean_ctor_set(x_143, 0, x_120); -lean_ctor_set(x_143, 1, x_131); -x_144 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_144, 0, x_119); -lean_ctor_set(x_144, 1, x_143); -x_145 = 1; -x_146 = lean_usize_add(x_4, x_145); -x_4 = x_146; -x_5 = x_144; -x_10 = x_126; +size_t x_37; size_t x_38; +lean_free_object(x_21); +lean_ctor_set(x_15, 1, x_28); +x_37 = 1; +x_38 = lean_usize_add(x_4, x_37); +x_4 = x_38; +x_10 = x_24; goto _start; } } else { -lean_object* x_148; lean_object* x_149; size_t x_150; size_t x_151; -lean_dec(x_127); -lean_dec(x_125); -if (lean_is_scalar(x_122)) { - x_148 = lean_alloc_ctor(0, 2, 0); -} else { - x_148 = x_122; -} -lean_ctor_set(x_148, 0, x_120); -lean_ctor_set(x_148, 1, x_131); -x_149 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_149, 0, x_119); -lean_ctor_set(x_149, 1, x_148); -x_150 = 1; -x_151 = lean_usize_add(x_4, x_150); -x_4 = x_151; -x_5 = x_149; -x_10 = x_126; +size_t x_40; size_t x_41; +lean_free_object(x_21); +lean_dec(x_23); +lean_ctor_set(x_15, 1, x_28); +x_40 = 1; +x_41 = lean_usize_add(x_4, x_40); +x_4 = x_41; +x_10 = x_24; goto _start; } } else { -uint8_t x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; size_t x_157; size_t x_158; -lean_dec(x_127); -lean_dec(x_125); -lean_dec(x_120); +uint8_t x_43; lean_object* x_44; size_t x_45; size_t x_46; +lean_free_object(x_21); +lean_dec(x_23); +lean_dec(x_18); lean_dec(x_13); -x_153 = 1; -x_154 = lean_box(x_153); -if (lean_is_scalar(x_122)) { - x_155 = lean_alloc_ctor(0, 2, 0); -} else { - x_155 = x_122; -} -lean_ctor_set(x_155, 0, x_154); -lean_ctor_set(x_155, 1, x_121); -x_156 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_156, 0, x_119); -lean_ctor_set(x_156, 1, x_155); -x_157 = 1; -x_158 = lean_usize_add(x_4, x_157); -x_4 = x_158; -x_5 = x_156; -x_10 = x_126; +x_43 = 1; +x_44 = lean_box(x_43); +lean_ctor_set(x_15, 0, x_44); +x_45 = 1; +x_46 = lean_usize_add(x_4, x_45); +x_4 = x_46; +x_10 = x_24; goto _start; } } else { -lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; -lean_dec(x_122); -lean_dec(x_121); -lean_dec(x_120); -lean_dec(x_119); -lean_dec(x_13); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_160 = lean_ctor_get(x_124, 0); -lean_inc(x_160); -x_161 = lean_ctor_get(x_124, 1); -lean_inc(x_161); -if (lean_is_exclusive(x_124)) { - lean_ctor_release(x_124, 0); - lean_ctor_release(x_124, 1); - x_162 = x_124; -} else { - lean_dec_ref(x_124); - x_162 = lean_box(0); -} -if (lean_is_scalar(x_162)) { - x_163 = lean_alloc_ctor(1, 2, 0); -} else { - x_163 = x_162; -} -lean_ctor_set(x_163, 0, x_160); -lean_ctor_set(x_163, 1, x_161); -return x_163; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_guessLex___spec__12___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = l_Array_eraseIdx___rarg(x_1, x_2); -x_13 = lean_array_push(x_3, x_4); -x_14 = l_Lean_Elab_WF_GuessLex_solve_go___at_Lean_Elab_WF_guessLex___spec__10(x_12, x_5, x_13, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_14) == 0) +lean_object* x_48; lean_object* x_49; uint8_t x_50; uint8_t x_51; uint8_t x_52; +x_48 = lean_ctor_get(x_21, 0); +x_49 = lean_ctor_get(x_21, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_21); +x_50 = 0; +x_51 = lean_unbox(x_48); +x_52 = l_Lean_Elab_WF_GuessLex_instDecidableEqGuessLexRel(x_51, x_50); +if (x_52 == 0) { -uint8_t x_15; -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) +lean_object* x_53; uint8_t x_54; uint8_t x_55; uint8_t x_56; +x_53 = lean_array_push(x_19, x_13); +x_54 = 2; +x_55 = lean_unbox(x_48); +x_56 = l_Lean_Elab_WF_GuessLex_instDecidableEqGuessLexRel(x_55, x_54); +if (x_56 == 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_14, 0); -x_17 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_17, 0, x_16); -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); -x_20 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_14, 0, x_20); -return x_14; +uint8_t x_57; uint8_t x_58; uint8_t x_59; +x_57 = 1; +x_58 = lean_unbox(x_48); +lean_dec(x_48); +x_59 = l_Lean_Elab_WF_GuessLex_instDecidableEqGuessLexRel(x_58, x_57); +if (x_59 == 0) +{ +uint8_t x_60; lean_object* x_61; lean_object* x_62; +lean_dec(x_17); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +lean_ctor_set(x_15, 1, x_53); +x_60 = 0; +x_61 = lean_box(x_60); +lean_ctor_set(x_5, 0, x_61); +x_62 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_62, 0, x_5); +lean_ctor_set(x_62, 1, x_49); +return x_62; } 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; -x_21 = lean_ctor_get(x_14, 0); -x_22 = lean_ctor_get(x_14, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_14); -x_23 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_23, 0, x_21); -x_24 = lean_box(0); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_23); -lean_ctor_set(x_25, 1, x_24); -x_26 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_26, 0, 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; +size_t x_63; size_t x_64; +lean_ctor_set(x_15, 1, x_53); +x_63 = 1; +x_64 = lean_usize_add(x_4, x_63); +x_4 = x_64; +x_10 = x_49; +goto _start; } } else { -uint8_t x_28; -x_28 = !lean_is_exclusive(x_14); -if (x_28 == 0) -{ -return x_14; +size_t x_66; size_t x_67; +lean_dec(x_48); +lean_ctor_set(x_15, 1, x_53); +x_66 = 1; +x_67 = lean_usize_add(x_4, x_66); +x_4 = x_67; +x_10 = x_49; +goto _start; +} } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_14, 0); -x_30 = lean_ctor_get(x_14, 1); -lean_inc(x_30); -lean_inc(x_29); -lean_dec(x_14); -x_31 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_31, 0, x_29); -lean_ctor_set(x_31, 1, x_30); -return x_31; -} +uint8_t x_69; lean_object* x_70; size_t x_71; size_t x_72; +lean_dec(x_48); +lean_dec(x_18); +lean_dec(x_13); +x_69 = 1; +x_70 = lean_box(x_69); +lean_ctor_set(x_15, 0, x_70); +x_71 = 1; +x_72 = lean_usize_add(x_4, x_71); +x_4 = x_72; +x_10 = x_49; +goto _start; } } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_guessLex___spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, 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; -x_19 = lean_nat_dec_lt(x_11, x_8); -if (x_19 == 0) +else { -lean_object* x_20; +uint8_t x_74; +lean_free_object(x_15); +lean_dec(x_19); +lean_dec(x_18); +lean_free_object(x_5); lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_11); -lean_dec(x_10); +lean_dec(x_13); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_3); lean_dec(x_1); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_13); -lean_ctor_set(x_20, 1, x_18); -return x_20; +x_74 = !lean_is_exclusive(x_21); +if (x_74 == 0) +{ +return x_21; } else { -lean_object* x_21; uint8_t x_22; -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; size_t x_27; size_t x_28; lean_object* x_29; lean_object* x_30; -lean_dec(x_13); -x_23 = lean_unsigned_to_nat(1u); -x_24 = lean_nat_sub(x_10, x_23); -lean_dec(x_10); -x_25 = lean_array_fget(x_1, x_11); -x_26 = lean_array_get_size(x_2); -x_27 = lean_usize_of_nat(x_26); -lean_dec(x_26); -x_28 = 0; -x_29 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2___rarg___closed__2; -lean_inc(x_17); -lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_25); -x_30 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__11(x_25, x_2, x_27, x_28, x_29, x_14, x_15, x_16, x_17, x_18); -if (lean_obj_tag(x_30) == 0) -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; -x_31 = lean_ctor_get(x_30, 0); -lean_inc(x_31); -x_32 = lean_ctor_get(x_31, 1); -lean_inc(x_32); -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_35; lean_object* x_36; -lean_dec(x_32); -lean_dec(x_31); -lean_dec(x_25); -x_35 = lean_ctor_get(x_30, 1); -lean_inc(x_35); -lean_dec(x_30); -x_36 = lean_nat_add(x_11, x_9); -lean_dec(x_11); -lean_inc(x_6); -{ -lean_object* _tmp_9 = x_24; -lean_object* _tmp_10 = x_36; -lean_object* _tmp_11 = lean_box(0); -lean_object* _tmp_12 = x_6; -lean_object* _tmp_17 = x_35; -x_10 = _tmp_9; -x_11 = _tmp_10; -x_12 = _tmp_11; -x_13 = _tmp_12; -x_18 = _tmp_17; +lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_75 = lean_ctor_get(x_21, 0); +x_76 = lean_ctor_get(x_21, 1); +lean_inc(x_76); +lean_inc(x_75); +lean_dec(x_21); +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; +} } -goto _start; } else { -lean_object* x_38; uint8_t x_39; -x_38 = lean_ctor_get(x_31, 0); -lean_inc(x_38); -lean_dec(x_31); -x_39 = lean_unbox(x_38); -lean_dec(x_38); -if (x_39 == 0) -{ -lean_object* x_40; lean_object* x_41; -lean_dec(x_32); -lean_dec(x_25); -x_40 = lean_ctor_get(x_30, 1); -lean_inc(x_40); -lean_dec(x_30); -x_41 = lean_nat_add(x_11, x_9); -lean_dec(x_11); +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_78 = lean_ctor_get(x_5, 0); +x_79 = lean_ctor_get(x_15, 0); +x_80 = lean_ctor_get(x_15, 1); +lean_inc(x_80); +lean_inc(x_79); +lean_dec(x_15); +x_81 = lean_array_uget(x_2, x_4); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); lean_inc(x_6); +lean_inc(x_1); +x_82 = lean_apply_6(x_81, x_1, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_82) == 0) { -lean_object* _tmp_9 = x_24; -lean_object* _tmp_10 = x_41; -lean_object* _tmp_11 = lean_box(0); -lean_object* _tmp_12 = x_6; -lean_object* _tmp_17 = x_40; -x_10 = _tmp_9; -x_11 = _tmp_10; -x_12 = _tmp_11; -x_13 = _tmp_12; -x_18 = _tmp_17; -} -goto _start; +lean_object* x_83; lean_object* x_84; lean_object* x_85; uint8_t x_86; uint8_t x_87; uint8_t x_88; +x_83 = lean_ctor_get(x_82, 0); +lean_inc(x_83); +x_84 = lean_ctor_get(x_82, 1); +lean_inc(x_84); +if (lean_is_exclusive(x_82)) { + lean_ctor_release(x_82, 0); + lean_ctor_release(x_82, 1); + x_85 = x_82; +} else { + lean_dec_ref(x_82); + x_85 = lean_box(0); } -else -{ -lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_43 = lean_ctor_get(x_30, 1); -lean_inc(x_43); -lean_dec(x_30); -x_44 = lean_ctor_get(x_32, 1); -lean_inc(x_44); -lean_dec(x_32); -x_45 = lean_box(0); -lean_inc(x_17); -lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_3); -lean_inc(x_1); -x_46 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_guessLex___spec__12___lambda__1(x_1, x_11, x_3, x_25, x_44, x_45, x_14, x_15, x_16, x_17, x_43); -lean_dec(x_44); -if (lean_obj_tag(x_46) == 0) +x_86 = 0; +x_87 = lean_unbox(x_83); +x_88 = l_Lean_Elab_WF_GuessLex_instDecidableEqGuessLexRel(x_87, x_86); +if (x_88 == 0) { -lean_object* x_47; -x_47 = lean_ctor_get(x_46, 0); -lean_inc(x_47); -if (lean_obj_tag(x_47) == 0) +lean_object* x_89; uint8_t x_90; uint8_t x_91; uint8_t x_92; +x_89 = lean_array_push(x_80, x_13); +x_90 = 2; +x_91 = lean_unbox(x_83); +x_92 = l_Lean_Elab_WF_GuessLex_instDecidableEqGuessLexRel(x_91, x_90); +if (x_92 == 0) { -uint8_t x_48; -lean_dec(x_24); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_11); +uint8_t x_93; uint8_t x_94; uint8_t x_95; +x_93 = 1; +x_94 = lean_unbox(x_83); +lean_dec(x_83); +x_95 = l_Lean_Elab_WF_GuessLex_instDecidableEqGuessLexRel(x_94, x_93); +if (x_95 == 0) +{ +lean_object* x_96; uint8_t x_97; lean_object* x_98; lean_object* x_99; +lean_dec(x_78); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_3); lean_dec(x_1); -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; +x_96 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_96, 0, x_79); +lean_ctor_set(x_96, 1, x_89); +x_97 = 0; +x_98 = lean_box(x_97); +lean_ctor_set(x_5, 1, x_96); +lean_ctor_set(x_5, 0, x_98); +if (lean_is_scalar(x_85)) { + x_99 = lean_alloc_ctor(0, 2, 0); +} else { + x_99 = x_85; +} +lean_ctor_set(x_99, 0, x_5); +lean_ctor_set(x_99, 1, x_84); +return x_99; } 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_object* x_100; size_t x_101; size_t x_102; +lean_dec(x_85); +x_100 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_100, 0, x_79); +lean_ctor_set(x_100, 1, x_89); +lean_ctor_set(x_5, 1, x_100); +x_101 = 1; +x_102 = lean_usize_add(x_4, x_101); +x_4 = x_102; +x_10 = x_84; +goto _start; } } else { -lean_object* x_54; lean_object* x_55; lean_object* x_56; -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 = lean_nat_add(x_11, x_9); -lean_dec(x_11); -x_10 = x_24; -x_11 = x_56; -x_12 = lean_box(0); -x_13 = x_55; -x_18 = x_54; +lean_object* x_104; size_t x_105; size_t x_106; +lean_dec(x_85); +lean_dec(x_83); +x_104 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_104, 0, x_79); +lean_ctor_set(x_104, 1, x_89); +lean_ctor_set(x_5, 1, x_104); +x_105 = 1; +x_106 = lean_usize_add(x_4, x_105); +x_4 = x_106; +x_10 = x_84; goto _start; } } else { -uint8_t x_58; -lean_dec(x_24); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_11); -lean_dec(x_6); -lean_dec(x_3); -lean_dec(x_1); -x_58 = !lean_is_exclusive(x_46); -if (x_58 == 0) -{ -return x_46; +uint8_t x_108; lean_object* x_109; lean_object* x_110; size_t x_111; size_t x_112; +lean_dec(x_85); +lean_dec(x_83); +lean_dec(x_79); +lean_dec(x_13); +x_108 = 1; +x_109 = lean_box(x_108); +x_110 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_110, 0, x_109); +lean_ctor_set(x_110, 1, x_80); +lean_ctor_set(x_5, 1, x_110); +x_111 = 1; +x_112 = lean_usize_add(x_4, x_111); +x_4 = x_112; +x_10 = x_84; +goto _start; +} } else { -lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_59 = lean_ctor_get(x_46, 0); -x_60 = lean_ctor_get(x_46, 1); -lean_inc(x_60); -lean_inc(x_59); -lean_dec(x_46); -x_61 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_61, 0, x_59); -lean_ctor_set(x_61, 1, x_60); -return x_61; +lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; +lean_dec(x_80); +lean_dec(x_79); +lean_free_object(x_5); +lean_dec(x_78); +lean_dec(x_13); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_114 = lean_ctor_get(x_82, 0); +lean_inc(x_114); +x_115 = lean_ctor_get(x_82, 1); +lean_inc(x_115); +if (lean_is_exclusive(x_82)) { + lean_ctor_release(x_82, 0); + lean_ctor_release(x_82, 1); + x_116 = x_82; +} else { + lean_dec_ref(x_82); + x_116 = lean_box(0); } +if (lean_is_scalar(x_116)) { + x_117 = lean_alloc_ctor(1, 2, 0); +} else { + x_117 = x_116; } +lean_ctor_set(x_117, 0, x_114); +lean_ctor_set(x_117, 1, x_115); +return x_117; } } } else { -uint8_t x_62; -lean_dec(x_25); -lean_dec(x_24); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_11); -lean_dec(x_6); -lean_dec(x_3); -lean_dec(x_1); -x_62 = !lean_is_exclusive(x_30); -if (x_62 == 0) -{ -return x_30; +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_5, 1); +x_119 = lean_ctor_get(x_5, 0); +lean_inc(x_118); +lean_inc(x_119); +lean_dec(x_5); +x_120 = lean_ctor_get(x_118, 0); +lean_inc(x_120); +x_121 = lean_ctor_get(x_118, 1); +lean_inc(x_121); +if (lean_is_exclusive(x_118)) { + lean_ctor_release(x_118, 0); + lean_ctor_release(x_118, 1); + x_122 = x_118; +} else { + lean_dec_ref(x_118); + x_122 = lean_box(0); } -else +x_123 = lean_array_uget(x_2, x_4); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_1); +x_124 = lean_apply_6(x_123, x_1, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_124) == 0) { -lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_63 = lean_ctor_get(x_30, 0); -x_64 = lean_ctor_get(x_30, 1); -lean_inc(x_64); -lean_inc(x_63); -lean_dec(x_30); -x_65 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_65, 0, x_63); -lean_ctor_set(x_65, 1, x_64); -return x_65; -} -} +lean_object* x_125; lean_object* x_126; lean_object* x_127; uint8_t x_128; uint8_t x_129; uint8_t x_130; +x_125 = lean_ctor_get(x_124, 0); +lean_inc(x_125); +x_126 = lean_ctor_get(x_124, 1); +lean_inc(x_126); +if (lean_is_exclusive(x_124)) { + lean_ctor_release(x_124, 0); + lean_ctor_release(x_124, 1); + x_127 = x_124; +} else { + lean_dec_ref(x_124); + x_127 = lean_box(0); } -else +x_128 = 0; +x_129 = lean_unbox(x_125); +x_130 = l_Lean_Elab_WF_GuessLex_instDecidableEqGuessLexRel(x_129, x_128); +if (x_130 == 0) +{ +lean_object* x_131; uint8_t x_132; uint8_t x_133; uint8_t x_134; +x_131 = lean_array_push(x_121, x_13); +x_132 = 2; +x_133 = lean_unbox(x_125); +x_134 = l_Lean_Elab_WF_GuessLex_instDecidableEqGuessLexRel(x_133, x_132); +if (x_134 == 0) +{ +uint8_t x_135; uint8_t x_136; uint8_t x_137; +x_135 = 1; +x_136 = lean_unbox(x_125); +lean_dec(x_125); +x_137 = l_Lean_Elab_WF_GuessLex_instDecidableEqGuessLexRel(x_136, x_135); +if (x_137 == 0) { -lean_object* x_66; -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_11); -lean_dec(x_10); +lean_object* x_138; uint8_t x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; +lean_dec(x_119); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_3); lean_dec(x_1); -x_66 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_66, 0, x_13); -lean_ctor_set(x_66, 1, x_18); -return x_66; -} -} +if (lean_is_scalar(x_122)) { + x_138 = lean_alloc_ctor(0, 2, 0); +} else { + x_138 = x_122; } +lean_ctor_set(x_138, 0, x_120); +lean_ctor_set(x_138, 1, x_131); +x_139 = 0; +x_140 = lean_box(x_139); +x_141 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_141, 0, x_140); +lean_ctor_set(x_141, 1, x_138); +if (lean_is_scalar(x_127)) { + x_142 = lean_alloc_ctor(0, 2, 0); +} else { + x_142 = x_127; } -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve_go___at_Lean_Elab_WF_guessLex___spec__10___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_10 = lean_array_get_size(x_1); -x_11 = lean_unsigned_to_nat(0u); -x_12 = lean_unsigned_to_nat(1u); -lean_inc(x_10); -x_13 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_13, 0, x_11); -lean_ctor_set(x_13, 1, x_10); -lean_ctor_set(x_13, 2, x_12); -x_14 = lean_box(0); -x_15 = l_Lean_Elab_WF_GuessLex_evalRecCall___lambda__2___closed__4; -lean_inc(x_10); -x_16 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_guessLex___spec__12(x_1, x_2, x_3, x_10, x_13, x_15, x_11, x_10, x_12, x_10, x_11, lean_box(0), x_15, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_13); -lean_dec(x_10); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; lean_object* x_18; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -lean_dec(x_17); -if (lean_obj_tag(x_18) == 0) -{ -uint8_t x_19; -x_19 = !lean_is_exclusive(x_16); -if (x_19 == 0) -{ -lean_object* x_20; -x_20 = lean_ctor_get(x_16, 0); -lean_dec(x_20); -lean_ctor_set(x_16, 0, x_14); -return x_16; +lean_ctor_set(x_142, 0, x_141); +lean_ctor_set(x_142, 1, x_126); +return x_142; } else { -lean_object* x_21; lean_object* x_22; -x_21 = lean_ctor_get(x_16, 1); -lean_inc(x_21); -lean_dec(x_16); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_14); -lean_ctor_set(x_22, 1, x_21); -return x_22; +lean_object* x_143; lean_object* x_144; size_t x_145; size_t x_146; +lean_dec(x_127); +if (lean_is_scalar(x_122)) { + x_143 = lean_alloc_ctor(0, 2, 0); +} else { + x_143 = x_122; } +lean_ctor_set(x_143, 0, x_120); +lean_ctor_set(x_143, 1, x_131); +x_144 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_144, 0, x_119); +lean_ctor_set(x_144, 1, x_143); +x_145 = 1; +x_146 = lean_usize_add(x_4, x_145); +x_4 = x_146; +x_5 = x_144; +x_10 = x_126; +goto _start; } -else -{ -uint8_t x_23; -x_23 = !lean_is_exclusive(x_16); -if (x_23 == 0) -{ -lean_object* x_24; lean_object* x_25; -x_24 = lean_ctor_get(x_16, 0); -lean_dec(x_24); -x_25 = lean_ctor_get(x_18, 0); -lean_inc(x_25); -lean_dec(x_18); -lean_ctor_set(x_16, 0, x_25); -return x_16; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_16, 1); -lean_inc(x_26); -lean_dec(x_16); -x_27 = lean_ctor_get(x_18, 0); -lean_inc(x_27); -lean_dec(x_18); -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; -} +lean_object* x_148; lean_object* x_149; size_t x_150; size_t x_151; +lean_dec(x_127); +lean_dec(x_125); +if (lean_is_scalar(x_122)) { + x_148 = lean_alloc_ctor(0, 2, 0); +} else { + x_148 = x_122; } +lean_ctor_set(x_148, 0, x_120); +lean_ctor_set(x_148, 1, x_131); +x_149 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_149, 0, x_119); +lean_ctor_set(x_149, 1, x_148); +x_150 = 1; +x_151 = lean_usize_add(x_4, x_150); +x_4 = x_151; +x_5 = x_149; +x_10 = x_126; +goto _start; } -else -{ -uint8_t x_29; -x_29 = !lean_is_exclusive(x_16); -if (x_29 == 0) -{ -return x_16; } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_16, 0); -x_31 = lean_ctor_get(x_16, 1); -lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_16); -x_32 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -return x_32; -} -} +uint8_t x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; size_t x_157; size_t x_158; +lean_dec(x_127); +lean_dec(x_125); +lean_dec(x_120); +lean_dec(x_13); +x_153 = 1; +x_154 = lean_box(x_153); +if (lean_is_scalar(x_122)) { + x_155 = lean_alloc_ctor(0, 2, 0); +} else { + x_155 = x_122; } +lean_ctor_set(x_155, 0, x_154); +lean_ctor_set(x_155, 1, x_121); +x_156 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_156, 0, x_119); +lean_ctor_set(x_156, 1, x_155); +x_157 = 1; +x_158 = lean_usize_add(x_4, x_157); +x_4 = x_158; +x_5 = x_156; +x_10 = x_126; +goto _start; } -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve_go___at_Lean_Elab_WF_guessLex___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) { -_start: -{ -uint8_t x_9; -x_9 = l_Array_isEmpty___rarg(x_2); -if (x_9 == 0) -{ -lean_object* x_10; lean_object* x_11; -x_10 = lean_box(0); -x_11 = l_Lean_Elab_WF_GuessLex_solve_go___at_Lean_Elab_WF_guessLex___spec__10___lambda__1(x_1, x_2, x_3, x_10, x_4, x_5, x_6, x_7, x_8); -return x_11; } else { -lean_object* x_12; lean_object* x_13; +lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; +lean_dec(x_122); +lean_dec(x_121); +lean_dec(x_120); +lean_dec(x_119); +lean_dec(x_13); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); lean_dec(x_1); -x_12 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_12, 0, x_3); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_12); -lean_ctor_set(x_13, 1, x_8); -return x_13; -} -} +x_160 = lean_ctor_get(x_124, 0); +lean_inc(x_160); +x_161 = lean_ctor_get(x_124, 1); +lean_inc(x_161); +if (lean_is_exclusive(x_124)) { + lean_ctor_release(x_124, 0); + lean_ctor_release(x_124, 1); + x_162 = x_124; +} else { + lean_dec_ref(x_124); + x_162 = lean_box(0); } -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve___at_Lean_Elab_WF_guessLex___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; -x_8 = l_Lean_Elab_WF_GuessLex_naryVarNames___closed__1; -x_9 = l_Lean_Elab_WF_GuessLex_solve_go___at_Lean_Elab_WF_guessLex___spec__10(x_1, x_2, x_8, x_3, x_4, x_5, x_6, x_7); -return x_9; +if (lean_is_scalar(x_162)) { + x_163 = lean_alloc_ctor(1, 2, 0); +} else { + x_163 = x_162; } +lean_ctor_set(x_163, 0, x_160); +lean_ctor_set(x_163, 1, x_161); +return x_163; } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__13(size_t x_1, size_t x_2, lean_object* x_3) { -_start: -{ -uint8_t x_4; -x_4 = lean_usize_dec_lt(x_2, x_1); -if (x_4 == 0) -{ -return x_3; } -else -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; -x_5 = lean_array_uget(x_3, x_2); -x_6 = lean_unsigned_to_nat(0u); -x_7 = lean_array_uset(x_3, x_2, x_6); -x_8 = lean_ctor_get(x_5, 3); -lean_inc(x_8); -lean_dec(x_5); -x_9 = 1; -x_10 = lean_usize_add(x_2, x_9); -x_11 = lean_array_uset(x_7, x_2, x_8); -x_2 = x_10; -x_3 = x_11; -goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_WF_guessLex___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_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_guessLex___spec__11___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_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_12; lean_object* x_13; lean_object* x_14; +x_12 = l_Array_eraseIdx___rarg(x_1, x_2); +x_13 = lean_array_push(x_3, x_4); +x_14 = l_Lean_Elab_WF_GuessLex_solve_go___at_Lean_Elab_WF_guessLex___spec__9(x_12, x_5, x_13, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_14) == 0) { -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; +uint8_t x_15; +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_16 = lean_ctor_get(x_14, 0); +x_17 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_17, 0, x_16); +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); +x_20 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_14, 0, x_20); +return x_14; } 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_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; -} +lean_object* x_21; 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_21 = lean_ctor_get(x_14, 0); +x_22 = lean_ctor_get(x_14, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_14); +x_23 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_23, 0, x_21); +x_24 = lean_box(0); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +x_26 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_26, 0, 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; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___lambda__1___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Termination", 11); -return x_1; -} -} -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___lambda__1___closed__2() { -_start: +uint8_t x_28; +x_28 = !lean_is_exclusive(x_14); +if (x_28 == 0) { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("terminationBy", 13); -return x_1; -} +return x_14; } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___lambda__1___closed__3() { -_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_Lean_Elab_WF_GuessLex_initFn____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_9____closed__6; -x_2 = l_Lean_Elab_WF_GuessLex_mkTupleSyntax___closed__1; -x_3 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___lambda__1___closed__1; -x_4 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___lambda__1___closed__2; -x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); -return x_5; +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_14, 0); +x_30 = lean_ctor_get(x_14, 1); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_14); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +return x_31; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___lambda__1___closed__4() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Try this: ", 10); -return x_1; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_guessLex___spec__11(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_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_10; lean_object* x_11; -x_10 = lean_ctor_get(x_1, 6); -lean_inc(x_10); -lean_dec(x_1); -x_11 = lean_ctor_get(x_10, 1); -lean_inc(x_11); -lean_dec(x_10); -if (lean_obj_tag(x_11) == 0) +uint8_t x_19; +x_19 = lean_nat_dec_lt(x_11, x_8); +if (x_19 == 0) { -lean_object* x_12; lean_object* x_13; -lean_dec(x_8); -lean_dec(x_7); +lean_object* x_20; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_6); lean_dec(x_3); -x_12 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_12, 0, x_2); -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; +lean_dec(x_1); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_13); +lean_ctor_set(x_20, 1, x_18); +return x_20; } 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; -x_14 = lean_ctor_get(x_11, 0); +lean_object* x_21; uint8_t x_22; +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; size_t x_27; size_t x_28; lean_object* x_29; lean_object* x_30; +lean_dec(x_13); +x_23 = lean_unsigned_to_nat(1u); +x_24 = lean_nat_sub(x_10, x_23); +lean_dec(x_10); +x_25 = lean_array_fget(x_1, x_11); +x_26 = lean_array_get_size(x_2); +x_27 = lean_usize_of_nat(x_26); +lean_dec(x_26); +x_28 = 0; +x_29 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_GuessLex_solve_go___spec__2___rarg___closed__2; +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); lean_inc(x_14); +lean_inc(x_25); +x_30 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__10(x_25, x_2, x_27, x_28, x_29, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_30) == 0) +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +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_35; lean_object* x_36; +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_25); +x_35 = lean_ctor_get(x_30, 1); +lean_inc(x_35); +lean_dec(x_30); +x_36 = lean_nat_add(x_11, x_9); lean_dec(x_11); -lean_inc(x_7); -x_15 = l_Lean_Elab_WF_TerminationBy_unexpand(x_3, 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_ctor_get(x_15, 1); +lean_inc(x_6); +{ +lean_object* _tmp_9 = x_24; +lean_object* _tmp_10 = x_36; +lean_object* _tmp_11 = lean_box(0); +lean_object* _tmp_12 = x_6; +lean_object* _tmp_17 = x_35; +x_10 = _tmp_9; +x_11 = _tmp_10; +x_12 = _tmp_11; +x_13 = _tmp_12; +x_18 = _tmp_17; +} +goto _start; +} +else +{ +lean_object* x_38; uint8_t x_39; +x_38 = lean_ctor_get(x_31, 0); +lean_inc(x_38); +lean_dec(x_31); +x_39 = lean_unbox(x_38); +lean_dec(x_38); +if (x_39 == 0) +{ +lean_object* x_40; lean_object* x_41; +lean_dec(x_32); +lean_dec(x_25); +x_40 = lean_ctor_get(x_30, 1); +lean_inc(x_40); +lean_dec(x_30); +x_41 = lean_nat_add(x_11, x_9); +lean_dec(x_11); +lean_inc(x_6); +{ +lean_object* _tmp_9 = x_24; +lean_object* _tmp_10 = x_41; +lean_object* _tmp_11 = lean_box(0); +lean_object* _tmp_12 = x_6; +lean_object* _tmp_17 = x_40; +x_10 = _tmp_9; +x_11 = _tmp_10; +x_12 = _tmp_11; +x_13 = _tmp_12; +x_18 = _tmp_17; +} +goto _start; +} +else +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_43 = lean_ctor_get(x_30, 1); +lean_inc(x_43); +lean_dec(x_30); +x_44 = lean_ctor_get(x_32, 1); +lean_inc(x_44); +lean_dec(x_32); +x_45 = lean_box(0); lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_3); +lean_inc(x_1); +x_46 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_guessLex___spec__11___lambda__1(x_1, x_11, x_3, x_25, x_44, x_45, x_14, x_15, x_16, x_17, x_43); +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_24); +lean_dec(x_17); +lean_dec(x_16); lean_dec(x_15); -x_18 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___lambda__1___closed__3; -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_16); -x_20 = lean_box(0); -x_21 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_21, 0, x_19); -lean_ctor_set(x_21, 1, x_20); -lean_ctor_set(x_21, 2, x_20); -lean_ctor_set(x_21, 3, x_20); -lean_ctor_set(x_21, 4, x_20); -lean_ctor_set(x_21, 5, x_20); -x_22 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___lambda__1___closed__4; -x_23 = l_Lean_Meta_Tactic_TryThis_addSuggestion(x_14, x_21, x_20, x_22, x_20, x_5, x_6, x_7, x_8, x_17); -if (lean_obj_tag(x_23) == 0) +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_1); +x_48 = !lean_is_exclusive(x_46); +if (x_48 == 0) { -uint8_t x_24; -x_24 = !lean_is_exclusive(x_23); -if (x_24 == 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_25; lean_object* x_26; -x_25 = lean_ctor_get(x_23, 0); -lean_dec(x_25); -x_26 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_26, 0, x_2); -lean_ctor_set(x_23, 0, x_26); -return x_23; +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; +} } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_23, 1); -lean_inc(x_27); -lean_dec(x_23); -x_28 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_28, 0, x_2); -x_29 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_29, 0, x_28); -lean_ctor_set(x_29, 1, x_27); -return x_29; +lean_object* x_54; lean_object* x_55; lean_object* x_56; +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 = lean_nat_add(x_11, x_9); +lean_dec(x_11); +x_10 = x_24; +x_11 = x_56; +x_12 = lean_box(0); +x_13 = x_55; +x_18 = x_54; +goto _start; } } else { -uint8_t x_30; -lean_dec(x_2); -x_30 = !lean_is_exclusive(x_23); -if (x_30 == 0) +uint8_t x_58; +lean_dec(x_24); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_1); +x_58 = !lean_is_exclusive(x_46); +if (x_58 == 0) { -return x_23; +return x_46; } 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_object* x_59; lean_object* x_60; lean_object* x_61; +x_59 = lean_ctor_get(x_46, 0); +x_60 = lean_ctor_get(x_46, 1); +lean_inc(x_60); +lean_inc(x_59); +lean_dec(x_46); +x_61 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_61, 0, x_59); +lean_ctor_set(x_61, 1, x_60); +return x_61; } } } } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = l_Lean_Elab_WF_GuessLex_showInferredTerminationBy; -return x_1; -} -} -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___closed__2() { -_start: +uint8_t x_62; +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_1); +x_62 = !lean_is_exclusive(x_30); +if (x_62 == 0) { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Inferred termination argument:\n", 31); -return x_1; -} +return x_30; } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___closed__3() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___closed__2; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_63 = lean_ctor_get(x_30, 0); +x_64 = lean_ctor_get(x_30, 1); +lean_inc(x_64); +lean_inc(x_63); +lean_dec(x_30); +x_65 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_65, 0, x_63); +lean_ctor_set(x_65, 1, x_64); +return x_65; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15(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; -x_10 = lean_usize_dec_lt(x_3, x_2); -if (x_10 == 0) -{ -lean_object* x_11; -lean_dec(x_8); -lean_dec(x_7); -x_11 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_11, 0, x_4); -lean_ctor_set(x_11, 1, x_9); -return x_11; } else { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_12 = lean_array_uget(x_1, x_3); -x_13 = lean_ctor_get(x_4, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_4, 1); -lean_inc(x_14); -x_15 = lean_ctor_get(x_4, 2); -lean_inc(x_15); -x_16 = lean_nat_dec_lt(x_14, x_15); -if (x_16 == 0) -{ -lean_object* x_17; +lean_object* x_66; +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_8); -lean_dec(x_7); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_4); -lean_ctor_set(x_17, 1, x_9); -return x_17; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_1); +x_66 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_66, 0, x_13); +lean_ctor_set(x_66, 1, x_18); +return x_66; } -else -{ -uint8_t x_18; -x_18 = !lean_is_exclusive(x_4); -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_4, 2); -lean_dec(x_19); -x_20 = lean_ctor_get(x_4, 1); -lean_dec(x_20); -x_21 = lean_ctor_get(x_4, 0); -lean_dec(x_21); -x_22 = lean_array_fget(x_13, x_14); -x_23 = lean_unsigned_to_nat(1u); -x_24 = lean_nat_add(x_14, x_23); -lean_dec(x_14); -lean_ctor_set(x_4, 1, x_24); -x_25 = lean_ctor_get(x_7, 2); -lean_inc(x_25); -x_26 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___closed__1; -x_27 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_25, x_26); -lean_dec(x_25); -if (x_27 == 0) -{ -lean_object* x_28; lean_object* x_29; -x_28 = lean_box(0); -lean_inc(x_8); -lean_inc(x_7); -x_29 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___lambda__1(x_12, x_4, x_22, x_28, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_29) == 0) +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve_go___at_Lean_Elab_WF_guessLex___spec__9___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: { -lean_object* x_30; -x_30 = lean_ctor_get(x_29, 0); -lean_inc(x_30); -if (lean_obj_tag(x_30) == 0) +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_10 = lean_array_get_size(x_1); +x_11 = lean_unsigned_to_nat(0u); +x_12 = lean_unsigned_to_nat(1u); +lean_inc(x_10); +x_13 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_13, 0, x_11); +lean_ctor_set(x_13, 1, x_10); +lean_ctor_set(x_13, 2, x_12); +x_14 = lean_box(0); +x_15 = l_Lean_Elab_WF_GuessLex_evalRecCall___lambda__2___closed__4; +lean_inc(x_10); +x_16 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_guessLex___spec__11(x_1, x_2, x_3, x_10, x_13, x_15, x_11, x_10, x_12, x_10, x_11, lean_box(0), x_15, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_13); +lean_dec(x_10); +if (lean_obj_tag(x_16) == 0) { -uint8_t x_31; -lean_dec(x_8); -lean_dec(x_7); -x_31 = !lean_is_exclusive(x_29); -if (x_31 == 0) +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +lean_dec(x_17); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_32; lean_object* x_33; -x_32 = lean_ctor_get(x_29, 0); -lean_dec(x_32); -x_33 = lean_ctor_get(x_30, 0); -lean_inc(x_33); -lean_dec(x_30); -lean_ctor_set(x_29, 0, x_33); -return x_29; -} -else +uint8_t x_19; +x_19 = !lean_is_exclusive(x_16); +if (x_19 == 0) { -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_29, 1); -lean_inc(x_34); -lean_dec(x_29); -x_35 = lean_ctor_get(x_30, 0); -lean_inc(x_35); -lean_dec(x_30); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_34); -return x_36; -} +lean_object* x_20; +x_20 = lean_ctor_get(x_16, 0); +lean_dec(x_20); +lean_ctor_set(x_16, 0, x_14); +return x_16; } else { -lean_object* x_37; lean_object* x_38; size_t x_39; size_t x_40; -x_37 = lean_ctor_get(x_29, 1); -lean_inc(x_37); -lean_dec(x_29); -x_38 = lean_ctor_get(x_30, 0); -lean_inc(x_38); -lean_dec(x_30); -x_39 = 1; -x_40 = lean_usize_add(x_3, x_39); -x_3 = x_40; -x_4 = x_38; -x_9 = x_37; -goto _start; +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_16, 1); +lean_inc(x_21); +lean_dec(x_16); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_14); +lean_ctor_set(x_22, 1, x_21); +return x_22; } } else { -uint8_t x_42; -lean_dec(x_8); -lean_dec(x_7); -x_42 = !lean_is_exclusive(x_29); -if (x_42 == 0) +uint8_t x_23; +x_23 = !lean_is_exclusive(x_16); +if (x_23 == 0) { -return x_29; +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_16, 0); +lean_dec(x_24); +x_25 = lean_ctor_get(x_18, 0); +lean_inc(x_25); +lean_dec(x_18); +lean_ctor_set(x_16, 0, x_25); +return x_16; } else { -lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_43 = lean_ctor_get(x_29, 0); -x_44 = lean_ctor_get(x_29, 1); -lean_inc(x_44); -lean_inc(x_43); -lean_dec(x_29); -x_45 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_45, 0, x_43); -lean_ctor_set(x_45, 1, x_44); -return x_45; +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_16, 1); +lean_inc(x_26); +lean_dec(x_16); +x_27 = lean_ctor_get(x_18, 0); +lean_inc(x_27); +lean_dec(x_18); +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_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; -lean_inc(x_7); -lean_inc(x_22); -x_46 = l_Lean_Elab_WF_TerminationBy_unexpand(x_22, x_5, x_6, x_7, x_8, x_9); -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_ctor_get(x_12, 0); -lean_inc(x_49); -x_50 = l_Lean_MessageData_ofSyntax(x_47); -x_51 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___closed__3; -x_52 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_52, 0, x_51); -lean_ctor_set(x_52, 1, x_50); -x_53 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__21___rarg___lambda__2___closed__5; -x_54 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_54, 0, x_52); -lean_ctor_set(x_54, 1, x_53); -x_55 = 0; -x_56 = l_Lean_logAt___at_Lean_Meta_computeSynthOrder___spec__7(x_49, x_54, x_55, x_5, x_6, x_7, x_8, x_48); -lean_dec(x_49); -x_57 = lean_ctor_get(x_56, 0); -lean_inc(x_57); -x_58 = lean_ctor_get(x_56, 1); -lean_inc(x_58); -lean_dec(x_56); -lean_inc(x_8); -lean_inc(x_7); -x_59 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___lambda__1(x_12, x_4, x_22, x_57, x_5, x_6, x_7, x_8, x_58); -lean_dec(x_57); -if (lean_obj_tag(x_59) == 0) +uint8_t x_29; +x_29 = !lean_is_exclusive(x_16); +if (x_29 == 0) { -lean_object* x_60; -x_60 = lean_ctor_get(x_59, 0); -lean_inc(x_60); -if (lean_obj_tag(x_60) == 0) +return x_16; +} +else { -uint8_t x_61; -lean_dec(x_8); -lean_dec(x_7); -x_61 = !lean_is_exclusive(x_59); -if (x_61 == 0) +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_16, 0); +x_31 = lean_ctor_get(x_16, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_16); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +return x_32; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve_go___at_Lean_Elab_WF_guessLex___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_62; lean_object* x_63; -x_62 = lean_ctor_get(x_59, 0); -lean_dec(x_62); -x_63 = lean_ctor_get(x_60, 0); -lean_inc(x_63); -lean_dec(x_60); -lean_ctor_set(x_59, 0, x_63); -return x_59; +uint8_t x_9; +x_9 = l_Array_isEmpty___rarg(x_2); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_box(0); +x_11 = l_Lean_Elab_WF_GuessLex_solve_go___at_Lean_Elab_WF_guessLex___spec__9___lambda__1(x_1, x_2, x_3, x_10, x_4, x_5, x_6, x_7, x_8); +return x_11; } else { -lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_64 = lean_ctor_get(x_59, 1); -lean_inc(x_64); -lean_dec(x_59); -x_65 = lean_ctor_get(x_60, 0); -lean_inc(x_65); -lean_dec(x_60); -x_66 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_66, 0, x_65); -lean_ctor_set(x_66, 1, x_64); -return x_66; +lean_object* x_12; lean_object* x_13; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_12 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_12, 0, x_3); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_8); +return x_13; } } -else +} +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve___at_Lean_Elab_WF_guessLex___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) { +_start: { -lean_object* x_67; lean_object* x_68; size_t x_69; size_t x_70; -x_67 = lean_ctor_get(x_59, 1); -lean_inc(x_67); -lean_dec(x_59); -x_68 = lean_ctor_get(x_60, 0); -lean_inc(x_68); -lean_dec(x_60); -x_69 = 1; -x_70 = lean_usize_add(x_3, x_69); -x_3 = x_70; -x_4 = x_68; -x_9 = x_67; -goto _start; +lean_object* x_8; lean_object* x_9; +x_8 = l_Lean_Elab_WF_GuessLex_naryVarNames___closed__1; +x_9 = l_Lean_Elab_WF_GuessLex_solve_go___at_Lean_Elab_WF_guessLex___spec__9(x_1, x_2, x_8, x_3, x_4, x_5, x_6, x_7); +return x_9; } } -else +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__12(size_t x_1, size_t x_2, lean_object* x_3) { +_start: { -uint8_t x_72; -lean_dec(x_8); -lean_dec(x_7); -x_72 = !lean_is_exclusive(x_59); -if (x_72 == 0) +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) { -return x_59; +return x_3; } else { -lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_73 = lean_ctor_get(x_59, 0); -x_74 = lean_ctor_get(x_59, 1); -lean_inc(x_74); -lean_inc(x_73); -lean_dec(x_59); -x_75 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_75, 0, x_73); -lean_ctor_set(x_75, 1, x_74); -return x_75; +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = lean_ctor_get(x_5, 3); +lean_inc(x_8); +lean_dec(x_5); +x_9 = 1; +x_10 = lean_usize_add(x_2, x_9); +x_11 = lean_array_uset(x_7, x_2, x_8); +x_2 = x_10; +x_3 = x_11; +goto _start; } } } +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_WF_guessLex___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_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_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_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_dec(x_4); -x_76 = lean_array_fget(x_13, x_14); -x_77 = lean_unsigned_to_nat(1u); -x_78 = lean_nat_add(x_14, x_77); -lean_dec(x_14); -x_79 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_79, 0, x_13); -lean_ctor_set(x_79, 1, x_78); -lean_ctor_set(x_79, 2, x_15); -x_80 = lean_ctor_get(x_7, 2); -lean_inc(x_80); -x_81 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___closed__1; -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) -{ -lean_object* x_83; lean_object* x_84; -x_83 = lean_box(0); -lean_inc(x_8); +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_84 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___lambda__1(x_12, x_79, x_76, x_83, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_84) == 0) +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; +} +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__14(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: { -lean_object* x_85; -x_85 = lean_ctor_get(x_84, 0); -lean_inc(x_85); -if (lean_obj_tag(x_85) == 0) +uint8_t x_10; +x_10 = lean_usize_dec_lt(x_3, x_2); +if (x_10 == 0) { -lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; +lean_object* x_11; lean_dec(x_8); lean_dec(x_7); -x_86 = lean_ctor_get(x_84, 1); -lean_inc(x_86); -if (lean_is_exclusive(x_84)) { - lean_ctor_release(x_84, 0); - lean_ctor_release(x_84, 1); - x_87 = x_84; -} else { - lean_dec_ref(x_84); - x_87 = lean_box(0); -} -x_88 = lean_ctor_get(x_85, 0); -lean_inc(x_88); -lean_dec(x_85); -if (lean_is_scalar(x_87)) { - x_89 = lean_alloc_ctor(0, 2, 0); -} else { - x_89 = x_87; -} -lean_ctor_set(x_89, 0, x_88); -lean_ctor_set(x_89, 1, x_86); -return x_89; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_1); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_4); +lean_ctor_set(x_11, 1, x_9); +return x_11; } else { -lean_object* x_90; lean_object* x_91; size_t x_92; size_t x_93; -x_90 = lean_ctor_get(x_84, 1); -lean_inc(x_90); -lean_dec(x_84); -x_91 = lean_ctor_get(x_85, 0); -lean_inc(x_91); -lean_dec(x_85); -x_92 = 1; -x_93 = lean_usize_add(x_3, x_92); -x_3 = x_93; -x_4 = x_91; -x_9 = x_90; +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_12 = lean_array_uget(x_4, x_3); +x_13 = lean_unsigned_to_nat(0u); +x_14 = lean_array_uset(x_4, x_3, x_13); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_1); +x_15 = l_Lean_Elab_WF_GuessLex_getNatParams(x_1, x_12, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(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_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_20 = lean_array_uset(x_14, x_3, x_16); +x_3 = x_19; +x_4 = x_20; +x_9 = x_17; goto _start; } -} else { -lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; +uint8_t x_22; +lean_dec(x_14); lean_dec(x_8); lean_dec(x_7); -x_95 = lean_ctor_get(x_84, 0); -lean_inc(x_95); -x_96 = lean_ctor_get(x_84, 1); -lean_inc(x_96); -if (lean_is_exclusive(x_84)) { - lean_ctor_release(x_84, 0); - lean_ctor_release(x_84, 1); - x_97 = x_84; -} else { - lean_dec_ref(x_84); - x_97 = lean_box(0); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_1); +x_22 = !lean_is_exclusive(x_15); +if (x_22 == 0) +{ +return x_15; } -if (lean_is_scalar(x_97)) { - x_98 = lean_alloc_ctor(1, 2, 0); -} else { - x_98 = x_97; +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_15, 0); +x_24 = lean_ctor_get(x_15, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_15); +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_ctor_set(x_98, 0, x_95); -lean_ctor_set(x_98, 1, x_96); -return x_98; } } -else -{ -lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; uint8_t x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; -lean_inc(x_7); -lean_inc(x_76); -x_99 = l_Lean_Elab_WF_TerminationBy_unexpand(x_76, x_5, x_6, x_7, x_8, x_9); -x_100 = lean_ctor_get(x_99, 0); -lean_inc(x_100); -x_101 = lean_ctor_get(x_99, 1); -lean_inc(x_101); -lean_dec(x_99); -x_102 = lean_ctor_get(x_12, 0); -lean_inc(x_102); -x_103 = l_Lean_MessageData_ofSyntax(x_100); -x_104 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___closed__3; -x_105 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_105, 0, x_104); -lean_ctor_set(x_105, 1, x_103); -x_106 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__21___rarg___lambda__2___closed__5; -x_107 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_107, 0, x_105); -lean_ctor_set(x_107, 1, x_106); -x_108 = 0; -x_109 = l_Lean_logAt___at_Lean_Meta_computeSynthOrder___spec__7(x_102, x_107, x_108, x_5, x_6, x_7, x_8, x_101); -lean_dec(x_102); -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_8); -lean_inc(x_7); -x_112 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___lambda__1(x_12, x_79, x_76, x_110, x_5, x_6, x_7, x_8, x_111); -lean_dec(x_110); -if (lean_obj_tag(x_112) == 0) +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__15(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: { -lean_object* x_113; -x_113 = lean_ctor_get(x_112, 0); -lean_inc(x_113); -if (lean_obj_tag(x_113) == 0) +uint8_t x_10; +x_10 = lean_usize_dec_lt(x_3, x_2); +if (x_10 == 0) { -lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; +lean_object* x_11; lean_dec(x_8); lean_dec(x_7); -x_114 = lean_ctor_get(x_112, 1); -lean_inc(x_114); -if (lean_is_exclusive(x_112)) { - lean_ctor_release(x_112, 0); - lean_ctor_release(x_112, 1); - x_115 = x_112; -} else { - lean_dec_ref(x_112); - x_115 = lean_box(0); -} -x_116 = lean_ctor_get(x_113, 0); -lean_inc(x_116); -lean_dec(x_113); -if (lean_is_scalar(x_115)) { - x_117 = lean_alloc_ctor(0, 2, 0); -} else { - x_117 = x_115; -} -lean_ctor_set(x_117, 0, x_116); -lean_ctor_set(x_117, 1, x_114); -return x_117; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_1); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_4); +lean_ctor_set(x_11, 1, x_9); +return x_11; } else { -lean_object* x_118; lean_object* x_119; size_t x_120; size_t x_121; -x_118 = lean_ctor_get(x_112, 1); -lean_inc(x_118); -lean_dec(x_112); -x_119 = lean_ctor_get(x_113, 0); -lean_inc(x_119); -lean_dec(x_113); -x_120 = 1; -x_121 = lean_usize_add(x_3, x_120); -x_3 = x_121; -x_4 = x_119; -x_9 = x_118; +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_12 = lean_array_uget(x_4, x_3); +x_13 = lean_unsigned_to_nat(0u); +x_14 = lean_array_uset(x_4, x_3, x_13); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_1); +x_15 = l_Lean_Elab_WF_GuessLex_getSizeOfParams(x_1, x_12, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(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_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_20 = lean_array_uset(x_14, x_3, x_16); +x_3 = x_19; +x_4 = x_20; +x_9 = x_17; goto _start; } -} else { -lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; +uint8_t x_22; +lean_dec(x_14); lean_dec(x_8); lean_dec(x_7); -x_123 = lean_ctor_get(x_112, 0); -lean_inc(x_123); -x_124 = lean_ctor_get(x_112, 1); -lean_inc(x_124); -if (lean_is_exclusive(x_112)) { - lean_ctor_release(x_112, 0); - lean_ctor_release(x_112, 1); - x_125 = x_112; -} else { - lean_dec_ref(x_112); - x_125 = lean_box(0); -} -if (lean_is_scalar(x_125)) { - x_126 = lean_alloc_ctor(1, 2, 0); -} else { - x_126 = x_125; -} -lean_ctor_set(x_126, 0, x_123); -lean_ctor_set(x_126, 1, x_124); -return x_126; -} +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_1); +x_22 = !lean_is_exclusive(x_15); +if (x_22 == 0) +{ +return x_15; } +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_15, 0); +x_24 = lean_ctor_get(x_15, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_15); +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; } } } @@ -25817,7 +28365,7 @@ x_22 = lean_array_get_size(x_21); x_23 = lean_usize_of_nat(x_22); lean_dec(x_22); lean_inc(x_4); -x_24 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__7(x_4, x_5, x_6, x_23, x_6, x_21, x_12, x_13, x_14, x_15, x_19); +x_24 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__6(x_4, x_5, x_6, x_23, x_6, x_21, x_12, x_13, x_14, x_15, x_19); x_25 = lean_ctor_get(x_24, 0); lean_inc(x_25); x_26 = lean_ctor_get(x_24, 1); @@ -25827,12 +28375,12 @@ x_27 = lean_array_get_size(x_25); x_28 = lean_usize_of_nat(x_27); lean_dec(x_27); lean_inc(x_25); -x_29 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__8(x_28, x_6, x_25); +x_29 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__7(x_28, x_6, x_25); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -x_30 = l_Lean_Elab_WF_GuessLex_solve___at_Lean_Elab_WF_guessLex___spec__9(x_7, x_29, x_12, x_13, x_14, x_15, x_26); +x_30 = l_Lean_Elab_WF_GuessLex_solve___at_Lean_Elab_WF_guessLex___spec__8(x_7, x_29, x_12, x_13, x_14, x_15, x_26); lean_dec(x_29); if (lean_obj_tag(x_30) == 0) { @@ -25845,7 +28393,7 @@ lean_object* x_32; lean_object* x_33; lean_object* x_34; x_32 = lean_ctor_get(x_30, 1); lean_inc(x_32); lean_dec(x_30); -x_33 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__13(x_5, x_6, x_4); +x_33 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__12(x_5, x_6, x_4); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); @@ -25874,7 +28422,7 @@ x_42 = l_Lean_Elab_WF_guessLex___lambda__1___closed__7; 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_throwError___at_Lean_Elab_WF_guessLex___spec__14(x_43, x_12, x_13, x_14, x_15, x_36); +x_44 = l_Lean_throwError___at_Lean_Elab_WF_guessLex___spec__13(x_43, x_12, x_13, x_14, x_15, x_36); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); @@ -25910,7 +28458,7 @@ return x_48; } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_dec(x_25); x_49 = lean_ctor_get(x_30, 1); lean_inc(x_49); @@ -25922,61 +28470,86 @@ lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -x_51 = l_Lean_Elab_WF_GuessLex_buildTermWF(x_9, x_8, x_50, x_12, x_13, x_14, x_15, x_49); +x_51 = l_Lean_Elab_WF_GuessLex_buildTermWF(x_9, x_8, x_10, x_50, x_12, x_13, x_14, x_15, x_49); x_52 = lean_ctor_get(x_51, 0); lean_inc(x_52); x_53 = lean_ctor_get(x_51, 1); lean_inc(x_53); lean_dec(x_51); -x_54 = l_Lean_Elab_WF_GuessLex_trimTermWF(x_10, x_52); -x_55 = lean_array_get_size(x_54); -x_56 = lean_unsigned_to_nat(0u); -x_57 = l_Array_toSubarray___rarg(x_54, x_56, x_55); -x_58 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15(x_4, x_5, x_6, x_57, x_12, x_13, x_14, x_15, x_53); +x_54 = l_Lean_Elab_WF_GuessLex_reportWF(x_4, x_52, x_12, x_13, x_14, x_15, x_53); lean_dec(x_13); lean_dec(x_12); -lean_dec(x_4); -if (lean_obj_tag(x_58) == 0) +if (lean_obj_tag(x_54) == 0) { -uint8_t x_59; -x_59 = !lean_is_exclusive(x_58); -if (x_59 == 0) +uint8_t x_55; +x_55 = !lean_is_exclusive(x_54); +if (x_55 == 0) { -lean_object* x_60; -x_60 = lean_ctor_get(x_58, 0); -lean_dec(x_60); +lean_object* x_56; +x_56 = lean_ctor_get(x_54, 0); +lean_dec(x_56); +lean_ctor_set(x_54, 0, x_52); +return x_54; +} +else +{ +lean_object* x_57; lean_object* x_58; +x_57 = lean_ctor_get(x_54, 1); +lean_inc(x_57); +lean_dec(x_54); +x_58 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_58, 0, x_52); +lean_ctor_set(x_58, 1, x_57); return x_58; } +} else { -lean_object* x_61; lean_object* x_62; -x_61 = lean_ctor_get(x_58, 1); +uint8_t x_59; +lean_dec(x_52); +x_59 = !lean_is_exclusive(x_54); +if (x_59 == 0) +{ +return x_54; +} +else +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_60 = lean_ctor_get(x_54, 0); +x_61 = lean_ctor_get(x_54, 1); lean_inc(x_61); -lean_dec(x_58); -x_62 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_62, 0, x_52); +lean_inc(x_60); +lean_dec(x_54); +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 { uint8_t x_63; -lean_dec(x_52); -x_63 = !lean_is_exclusive(x_58); +lean_dec(x_25); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_4); +x_63 = !lean_is_exclusive(x_30); if (x_63 == 0) { -return x_58; +return x_30; } else { lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_64 = lean_ctor_get(x_58, 0); -x_65 = lean_ctor_get(x_58, 1); +x_64 = lean_ctor_get(x_30, 0); +x_65 = lean_ctor_get(x_30, 1); lean_inc(x_65); lean_inc(x_64); -lean_dec(x_58); +lean_dec(x_30); x_66 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_66, 0, x_64); lean_ctor_set(x_66, 1, x_65); @@ -25984,29 +28557,28 @@ return x_66; } } } -} else { uint8_t x_67; -lean_dec(x_25); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); +lean_dec(x_7); lean_dec(x_4); -x_67 = !lean_is_exclusive(x_30); +x_67 = !lean_is_exclusive(x_17); if (x_67 == 0) { -return x_30; +return x_17; } else { lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_68 = lean_ctor_get(x_30, 0); -x_69 = lean_ctor_get(x_30, 1); +x_68 = lean_ctor_get(x_17, 0); +x_69 = lean_ctor_get(x_17, 1); lean_inc(x_69); lean_inc(x_68); -lean_dec(x_30); +lean_dec(x_17); x_70 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_70, 0, x_68); lean_ctor_set(x_70, 1, x_69); @@ -26014,116 +28586,246 @@ return x_70; } } } +} +LEAN_EXPORT lean_object* l_Lean_Elab_WF_guessLex___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, size_t x_6, size_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +lean_inc(x_2); +x_16 = l_Lean_Elab_WF_GuessLex_generateMeasures(x_1, x_2, x_11, x_12, x_13, x_14, x_15); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = lean_array_get_size(x_17); +x_20 = lean_unsigned_to_nat(1u); +x_21 = lean_nat_dec_eq(x_19, x_20); +lean_dec(x_19); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +x_23 = l_Lean_Elab_WF_guessLex___lambda__1(x_3, x_4, x_2, x_5, x_6, x_7, x_17, x_8, x_9, x_10, x_22, x_11, x_12, x_13, x_14, x_18); +return x_23; +} else { -uint8_t x_71; -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_7); +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_dec(x_4); -x_71 = !lean_is_exclusive(x_17); -if (x_71 == 0) +lean_dec(x_3); +lean_dec(x_2); +x_24 = lean_unsigned_to_nat(0u); +x_25 = lean_array_fget(x_17, x_24); +lean_dec(x_17); +x_26 = l_Lean_Elab_WF_GuessLex_GuessLexRel_toNatRel___closed__20; +x_27 = lean_array_push(x_26, x_25); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +x_28 = l_Lean_Elab_WF_GuessLex_buildTermWF(x_9, x_8, x_10, x_27, x_11, x_12, x_13, x_14, x_18); +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_Lean_Elab_WF_GuessLex_reportWF(x_5, x_29, x_11, x_12, x_13, x_14, x_30); +lean_dec(x_12); +lean_dec(x_11); +if (lean_obj_tag(x_31) == 0) { -return x_17; +uint8_t x_32; +x_32 = !lean_is_exclusive(x_31); +if (x_32 == 0) +{ +lean_object* x_33; +x_33 = lean_ctor_get(x_31, 0); +lean_dec(x_33); +lean_ctor_set(x_31, 0, x_29); +return x_31; } else { -lean_object* x_72; lean_object* x_73; lean_object* x_74; -x_72 = lean_ctor_get(x_17, 0); -x_73 = lean_ctor_get(x_17, 1); -lean_inc(x_73); -lean_inc(x_72); +lean_object* x_34; lean_object* x_35; +x_34 = lean_ctor_get(x_31, 1); +lean_inc(x_34); +lean_dec(x_31); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_29); +lean_ctor_set(x_35, 1, x_34); +return x_35; +} +} +else +{ +uint8_t x_36; +lean_dec(x_29); +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; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_WF_guessLex___lambda__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_object* x_14, lean_object* x_15) { +_start: +{ +lean_object* x_16; +lean_dec(x_10); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_4); +lean_inc(x_1); +x_16 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__4(x_1, x_2, x_3, x_4, x_11, x_12, x_13, x_14, 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 = lean_unsigned_to_nat(1u); +x_20 = lean_nat_dec_eq(x_9, x_19); +lean_dec(x_9); +if (x_20 == 0) +{ +lean_object* x_21; +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_4); +lean_inc(x_1); +x_21 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__14(x_1, x_2, x_3, x_4, x_11, x_12, x_13, x_14, x_18); +if (lean_obj_tag(x_21) == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +x_24 = l_Lean_Elab_WF_guessLex___lambda__2(x_17, x_5, x_6, x_1, x_4, x_2, x_3, x_7, x_8, x_22, x_11, x_12, x_13, x_14, x_23); +lean_dec(x_22); +lean_dec(x_8); +lean_dec(x_7); +return x_24; +} +else +{ +uint8_t x_25; lean_dec(x_17); -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_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +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_1); +x_25 = !lean_is_exclusive(x_21); +if (x_25 == 0) +{ +return x_21; } +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_21, 0); +x_27 = lean_ctor_get(x_21, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_21); +x_28 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_WF_guessLex___lambda__2(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { -_start: +else { -lean_object* x_16; -lean_dec(x_10); +lean_object* x_29; lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_inc(x_4); lean_inc(x_1); -x_16 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__5(x_1, x_2, x_3, x_4, x_11, x_12, x_13, x_14, x_15); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); -lean_inc(x_18); -lean_dec(x_16); -lean_inc(x_5); -x_19 = l_Lean_Elab_WF_GuessLex_generateMeasures(x_17, x_5, x_11, x_12, x_13, x_14, x_18); -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -x_21 = lean_ctor_get(x_19, 1); -lean_inc(x_21); -lean_dec(x_19); -x_22 = lean_array_get_size(x_20); -x_23 = lean_unsigned_to_nat(1u); -x_24 = lean_nat_dec_eq(x_22, x_23); -lean_dec(x_22); -if (x_24 == 0) +x_29 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__15(x_1, x_2, x_3, x_4, x_11, x_12, x_13, x_14, x_18); +if (lean_obj_tag(x_29) == 0) { -lean_object* x_25; lean_object* x_26; -x_25 = lean_box(0); -x_26 = l_Lean_Elab_WF_guessLex___lambda__1(x_6, x_1, x_5, x_4, x_2, x_3, x_20, x_7, x_8, x_9, x_25, x_11, x_12, x_13, x_14, x_21); -lean_dec(x_9); +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_29, 1); +lean_inc(x_31); +lean_dec(x_29); +x_32 = l_Lean_Elab_WF_guessLex___lambda__2(x_17, x_5, x_6, x_1, x_4, x_2, x_3, x_7, x_8, x_30, x_11, x_12, x_13, x_14, x_31); +lean_dec(x_30); lean_dec(x_8); lean_dec(x_7); -return x_26; +return x_32; } 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; -lean_dec(x_9); +uint8_t x_33; +lean_dec(x_17); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +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_1); -x_27 = lean_unsigned_to_nat(0u); -x_28 = lean_array_fget(x_20, x_27); -lean_dec(x_20); -x_29 = l_Lean_Elab_WF_GuessLex_GuessLexRel_toNatRel___closed__20; -x_30 = lean_array_push(x_29, x_28); -x_31 = l_Lean_Elab_WF_GuessLex_buildTermWF(x_8, x_7, x_30, x_11, x_12, x_13, x_14, x_21); -lean_dec(x_7); -lean_dec(x_8); -x_32 = !lean_is_exclusive(x_31); -if (x_32 == 0) +x_33 = !lean_is_exclusive(x_29); +if (x_33 == 0) { -return x_31; +return x_29; } 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_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_29, 0); +x_35 = lean_ctor_get(x_29, 1); +lean_inc(x_35); 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_dec(x_29); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; +} } } } else { -uint8_t x_36; +uint8_t x_37; lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); @@ -26135,23 +28837,23 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_1); -x_36 = !lean_is_exclusive(x_16); -if (x_36 == 0) +x_37 = !lean_is_exclusive(x_16); +if (x_37 == 0) { return x_16; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_16, 0); -x_38 = lean_ctor_get(x_16, 1); +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_16, 0); +x_39 = lean_ctor_get(x_16, 1); +lean_inc(x_39); lean_inc(x_38); -lean_inc(x_37); lean_dec(x_16); -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; +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; } } } @@ -26176,101 +28878,98 @@ return x_2; LEAN_EXPORT lean_object* l_Lean_Elab_WF_guessLex(lean_object* x_1, lean_object* x_2, 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; size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; +lean_object* x_9; size_t x_10; size_t x_11; lean_object* x_12; x_9 = lean_array_get_size(x_1); x_10 = lean_usize_of_nat(x_9); -lean_dec(x_9); x_11 = 0; -lean_inc(x_1); -x_12 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__1(x_10, x_11, x_1); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_1); -x_13 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__2(x_10, x_11, x_1, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_13) == 0) +x_12 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__1(x_10, x_11, x_1, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_12) == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; size_t x_17; lean_object* x_18; -x_14 = lean_ctor_get(x_13, 0); +lean_object* x_13; lean_object* x_14; lean_object* x_15; size_t x_16; lean_object* x_17; +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); -x_16 = lean_array_get_size(x_14); -x_17 = lean_usize_of_nat(x_16); -lean_dec(x_16); +lean_dec(x_12); +x_15 = lean_array_get_size(x_13); +x_16 = lean_usize_of_nat(x_15); +lean_dec(x_15); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_14); +lean_inc(x_13); lean_inc(x_3); -x_18 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__3(x_3, x_17, x_11, x_14, x_4, x_5, x_6, x_7, x_15); -if (lean_obj_tag(x_18) == 0) +x_17 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__2(x_3, x_16, x_11, x_13, x_4, x_5, x_6, x_7, x_14); +if (lean_obj_tag(x_17) == 0) { -lean_object* x_19; lean_object* x_20; lean_object* x_21; size_t x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; -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_array_get_size(x_19); -x_22 = lean_usize_of_nat(x_21); -lean_dec(x_21); +lean_object* x_18; lean_object* x_19; lean_object* x_20; size_t x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t 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); -x_23 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__4(x_22, x_11, x_19); -x_24 = l_Lean_Elab_WF_GuessLex_withRecApps___rarg___closed__3; -x_25 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_24, x_4, x_5, x_6, x_7, x_20); -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; -x_28 = lean_ctor_get(x_25, 1); -lean_inc(x_28); +lean_dec(x_17); +x_20 = lean_array_get_size(x_18); +x_21 = lean_usize_of_nat(x_20); +lean_dec(x_20); +lean_inc(x_18); +x_22 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__3(x_21, x_11, x_18); +x_23 = l_Lean_Elab_WF_GuessLex_withRecApps___rarg___closed__3; +x_24 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_23, x_4, x_5, x_6, x_7, x_19); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_unbox(x_25); lean_dec(x_25); -x_29 = lean_box(0); -x_30 = l_Lean_Elab_WF_guessLex___lambda__2(x_3, x_10, x_11, x_1, x_23, x_2, x_19, x_14, x_12, x_29, x_4, x_5, x_6, x_7, x_28); -return x_30; +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_24, 1); +lean_inc(x_27); +lean_dec(x_24); +x_28 = lean_box(0); +x_29 = l_Lean_Elab_WF_guessLex___lambda__3(x_3, x_10, x_11, x_1, x_22, x_2, x_18, x_13, x_9, x_28, x_4, x_5, x_6, x_7, x_27); +return x_29; } 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; -x_31 = lean_ctor_get(x_25, 1); -lean_inc(x_31); -lean_dec(x_25); -lean_inc(x_19); -x_32 = lean_array_to_list(lean_box(0), x_19); -x_33 = lean_box(0); -x_34 = l_List_mapTR_loop___at_Lean_Elab_WF_guessLex___spec__16(x_32, x_33); -x_35 = l_Lean_MessageData_ofList(x_34); -lean_dec(x_34); -x_36 = l_Lean_Elab_WF_guessLex___closed__2; -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_Array_foldlMUnsafe_fold___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__21___rarg___lambda__2___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 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_24, x_39, x_4, x_5, x_6, x_7, x_31); -x_41 = lean_ctor_get(x_40, 0); +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_30 = lean_ctor_get(x_24, 1); +lean_inc(x_30); +lean_dec(x_24); +lean_inc(x_18); +x_31 = lean_array_to_list(lean_box(0), x_18); +x_32 = lean_box(0); +x_33 = l_List_mapTR_loop___at_Lean_Elab_WF_guessLex___spec__16(x_31, x_32); +x_34 = l_Lean_MessageData_ofList(x_33); +lean_dec(x_33); +x_35 = l_Lean_Elab_WF_guessLex___closed__2; +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_Array_foldlMUnsafe_fold___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__21___rarg___lambda__2___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_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_23, x_38, x_4, x_5, x_6, x_7, x_30); +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +x_41 = lean_ctor_get(x_39, 1); lean_inc(x_41); -x_42 = lean_ctor_get(x_40, 1); -lean_inc(x_42); -lean_dec(x_40); -x_43 = l_Lean_Elab_WF_guessLex___lambda__2(x_3, x_10, x_11, x_1, x_23, x_2, x_19, x_14, x_12, x_41, x_4, x_5, x_6, x_7, x_42); -return x_43; +lean_dec(x_39); +x_42 = l_Lean_Elab_WF_guessLex___lambda__3(x_3, x_10, x_11, x_1, x_22, x_2, x_18, x_13, x_9, x_40, x_4, x_5, x_6, x_7, x_41); +return x_42; } } else { -uint8_t x_44; -lean_dec(x_14); -lean_dec(x_12); +uint8_t x_43; +lean_dec(x_13); +lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); @@ -26278,30 +28977,30 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_44 = !lean_is_exclusive(x_18); -if (x_44 == 0) +x_43 = !lean_is_exclusive(x_17); +if (x_43 == 0) { -return x_18; +return x_17; } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_45 = lean_ctor_get(x_18, 0); -x_46 = lean_ctor_get(x_18, 1); -lean_inc(x_46); +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_17, 0); +x_45 = lean_ctor_get(x_17, 1); lean_inc(x_45); -lean_dec(x_18); -x_47 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_47, 0, x_45); -lean_ctor_set(x_47, 1, x_46); -return x_47; +lean_inc(x_44); +lean_dec(x_17); +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; } } } else { -uint8_t x_48; -lean_dec(x_12); +uint8_t x_47; +lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); @@ -26309,40 +29008,28 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_48 = !lean_is_exclusive(x_13); -if (x_48 == 0) +x_47 = !lean_is_exclusive(x_12); +if (x_47 == 0) { -return x_13; +return x_12; } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_49 = lean_ctor_get(x_13, 0); -x_50 = lean_ctor_get(x_13, 1); -lean_inc(x_50); +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_12, 0); +x_49 = lean_ctor_get(x_12, 1); lean_inc(x_49); -lean_dec(x_13); -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_inc(x_48); +lean_dec(x_12); +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_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -size_t x_4; size_t x_5; lean_object* x_6; -x_4 = lean_unbox_usize(x_1); -lean_dec(x_1); -x_5 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__1(x_4, x_5, x_3); -return x_6; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___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_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { size_t x_9; size_t x_10; lean_object* x_11; @@ -26350,11 +29037,11 @@ x_9 = lean_unbox_usize(x_1); lean_dec(x_1); x_10 = lean_unbox_usize(x_2); lean_dec(x_2); -x_11 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__2(x_9, x_10, x_3, x_4, x_5, x_6, x_7, x_8); +x_11 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__1(x_9, x_10, x_3, x_4, x_5, x_6, x_7, x_8); return x_11; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___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_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { size_t x_10; size_t x_11; lean_object* x_12; @@ -26362,11 +29049,11 @@ 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_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__3(x_1, x_10, x_11, x_4, x_5, x_6, x_7, x_8, x_9); +x_12 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__2(x_1, x_10, x_11, x_4, x_5, x_6, x_7, x_8, x_9); return x_12; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___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_Elab_WF_guessLex___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; @@ -26374,11 +29061,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_Elab_WF_guessLex___spec__4(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__3(x_4, x_5, x_3); return x_6; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___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_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { size_t x_10; size_t x_11; lean_object* x_12; @@ -26386,11 +29073,11 @@ 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_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__5(x_1, x_10, x_11, x_4, x_5, x_6, x_7, x_8, x_9); +x_12 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__4(x_1, x_10, x_11, x_4, x_5, x_6, x_7, x_8, x_9); return x_12; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { size_t x_4; size_t x_5; lean_object* x_6; @@ -26398,11 +29085,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_Elab_WF_guessLex___spec__6(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__5(x_4, x_5, x_3); return x_6; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___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_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { size_t x_12; size_t x_13; size_t x_14; size_t x_15; lean_object* x_16; @@ -26414,7 +29101,7 @@ x_14 = lean_unbox_usize(x_4); lean_dec(x_4); x_15 = lean_unbox_usize(x_5); lean_dec(x_5); -x_16 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__7(x_1, x_12, x_13, x_14, x_15, x_6, x_7, x_8, x_9, x_10, x_11); +x_16 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__6(x_1, x_12, x_13, x_14, x_15, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -26422,7 +29109,7 @@ lean_dec(x_7); return x_16; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__7___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; @@ -26430,11 +29117,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_Elab_WF_guessLex___spec__8(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__7(x_4, x_5, x_3); return x_6; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___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; @@ -26442,23 +29129,23 @@ 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_Elab_WF_guessLex___spec__11(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_Elab_WF_guessLex___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_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_guessLex___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, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_guessLex___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, 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_x27_loop___at_Lean_Elab_WF_guessLex___spec__12___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_12 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_guessLex___spec__11___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_2); return x_12; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_guessLex___spec__12___boxed(lean_object** _args) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_guessLex___spec__11___boxed(lean_object** _args) { lean_object* x_1 = _args[0]; lean_object* x_2 = _args[1]; lean_object* x_3 = _args[2]; @@ -26480,7 +29167,7 @@ lean_object* x_18 = _args[17]; _start: { lean_object* x_19; -x_19 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_guessLex___spec__12(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 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_WF_guessLex___spec__11(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -26490,35 +29177,35 @@ lean_dec(x_2); return x_19; } } -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve_go___at_Lean_Elab_WF_guessLex___spec__10___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve_go___at_Lean_Elab_WF_guessLex___spec__9___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_WF_GuessLex_solve_go___at_Lean_Elab_WF_guessLex___spec__10___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_10 = l_Lean_Elab_WF_GuessLex_solve_go___at_Lean_Elab_WF_guessLex___spec__9___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_2); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve_go___at_Lean_Elab_WF_guessLex___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_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve_go___at_Lean_Elab_WF_guessLex___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: { lean_object* x_9; -x_9 = l_Lean_Elab_WF_GuessLex_solve_go___at_Lean_Elab_WF_guessLex___spec__10(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_9 = l_Lean_Elab_WF_GuessLex_solve_go___at_Lean_Elab_WF_guessLex___spec__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_2); return x_9; } } -LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve___at_Lean_Elab_WF_guessLex___spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_solve___at_Lean_Elab_WF_guessLex___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) { _start: { lean_object* x_8; -x_8 = l_Lean_Elab_WF_GuessLex_solve___at_Lean_Elab_WF_guessLex___spec__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +x_8 = l_Lean_Elab_WF_GuessLex_solve___at_Lean_Elab_WF_guessLex___spec__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7); lean_dec(x_2); return x_8; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__13___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__12___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; @@ -26526,15 +29213,15 @@ x_4 = lean_unbox_usize(x_1); lean_dec(x_1); x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__13(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__12(x_4, x_5, x_3); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_WF_guessLex___spec__14___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_WF_guessLex___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) { _start: { lean_object* x_7; -x_7 = l_Lean_throwError___at_Lean_Elab_WF_guessLex___spec__14(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_Lean_throwError___at_Lean_Elab_WF_guessLex___spec__13(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); @@ -26542,18 +29229,19 @@ lean_dec(x_2); return x_7; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__14___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_10; -x_10 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_10; +size_t x_10; size_t x_11; lean_object* x_12; +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_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__14(x_1, x_10, x_11, x_4, x_5, x_6, x_7, x_8, x_9); +return x_12; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___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_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___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) { _start: { size_t x_10; size_t x_11; lean_object* x_12; @@ -26561,10 +29249,7 @@ 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_Elab_WF_guessLex___spec__15(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); +x_12 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__15(x_1, x_10, x_11, x_4, x_5, x_6, x_7, x_8, x_9); return x_12; } } @@ -26588,11 +29273,26 @@ LEAN_EXPORT lean_object* l_Lean_Elab_WF_guessLex___lambda__2___boxed(lean_object _start: { size_t x_16; size_t x_17; lean_object* x_18; +x_16 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_17 = lean_unbox_usize(x_7); +lean_dec(x_7); +x_18 = l_Lean_Elab_WF_guessLex___lambda__2(x_1, x_2, x_3, x_4, x_5, x_16, x_17, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_WF_guessLex___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: +{ +size_t x_16; size_t x_17; lean_object* x_18; x_16 = lean_unbox_usize(x_2); lean_dec(x_2); x_17 = lean_unbox_usize(x_3); lean_dec(x_3); -x_18 = l_Lean_Elab_WF_guessLex___lambda__2(x_1, x_16, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +x_18 = l_Lean_Elab_WF_guessLex___lambda__3(x_1, x_16, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); return x_18; } } @@ -27020,6 +29720,22 @@ l_Lean_Elab_WF_GuessLex_evalRecCall___closed__2 = _init_l_Lean_Elab_WF_GuessLex_ lean_mark_persistent(l_Lean_Elab_WF_GuessLex_evalRecCall___closed__2); l_Lean_Elab_WF_GuessLex_RecCallCache_prettyEntry___closed__1 = _init_l_Lean_Elab_WF_GuessLex_RecCallCache_prettyEntry___closed__1(); lean_mark_persistent(l_Lean_Elab_WF_GuessLex_RecCallCache_prettyEntry___closed__1); +l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__2___lambda__1___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__2___lambda__1___closed__1(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__2___lambda__1___closed__1); +l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__2___lambda__1___closed__2 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__2___lambda__1___closed__2(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__2___lambda__1___closed__2); +l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__2___lambda__1___closed__3 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__2___lambda__1___closed__3(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__2___lambda__1___closed__3); +l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__2___lambda__1___closed__4 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__2___lambda__1___closed__4(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getSizeOfParams___spec__2___lambda__1___closed__4); +l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getNatParams___spec__1___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getNatParams___spec__1___closed__1(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_getNatParams___spec__1___closed__1); +l_Lean_Elab_WF_GuessLex_getNatParams___closed__1 = _init_l_Lean_Elab_WF_GuessLex_getNatParams___closed__1(); +lean_mark_persistent(l_Lean_Elab_WF_GuessLex_getNatParams___closed__1); +l_Lean_Elab_WF_GuessLex_getNatParams___closed__2 = _init_l_Lean_Elab_WF_GuessLex_getNatParams___closed__2(); +lean_mark_persistent(l_Lean_Elab_WF_GuessLex_getNatParams___closed__2); +l_Lean_Elab_WF_GuessLex_getNatParams___closed__3 = _init_l_Lean_Elab_WF_GuessLex_getNatParams___closed__3(); +lean_mark_persistent(l_Lean_Elab_WF_GuessLex_getNatParams___closed__3); l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform___closed__1 = _init_l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform___closed__1(); lean_mark_persistent(l_Lean_Elab_WF_GuessLex_generateCombinations_x3f_goUniform___closed__1); l_Lean_Elab_WF_GuessLex_generateMeasures___closed__1 = _init_l_Lean_Elab_WF_GuessLex_generateMeasures___closed__1(); @@ -27111,20 +29827,20 @@ l_Lean_Elab_WF_GuessLex_explainFailure___closed__4 = _init_l_Lean_Elab_WF_GuessL lean_mark_persistent(l_Lean_Elab_WF_GuessLex_explainFailure___closed__4); l_Lean_Elab_WF_GuessLex_explainFailure___closed__5 = _init_l_Lean_Elab_WF_GuessLex_explainFailure___closed__5(); lean_mark_persistent(l_Lean_Elab_WF_GuessLex_explainFailure___closed__5); -l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___lambda__1___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___lambda__1___closed__1(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___lambda__1___closed__1); -l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___lambda__1___closed__2 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___lambda__1___closed__2(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___lambda__1___closed__2); -l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___lambda__1___closed__3 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___lambda__1___closed__3(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___lambda__1___closed__3); -l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___lambda__1___closed__4 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___lambda__1___closed__4(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___lambda__1___closed__4); -l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___closed__1(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___closed__1); -l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___closed__2 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___closed__2(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___closed__2); -l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___closed__3 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___closed__3(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_WF_guessLex___spec__15___closed__3); +l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___lambda__1___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___lambda__1___closed__1(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___lambda__1___closed__1); +l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___lambda__1___closed__2 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___lambda__1___closed__2(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___lambda__1___closed__2); +l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___lambda__1___closed__3 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___lambda__1___closed__3(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___lambda__1___closed__3); +l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___lambda__1___closed__4 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___lambda__1___closed__4(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___lambda__1___closed__4); +l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___closed__1(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___closed__1); +l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___closed__2 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___closed__2(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___closed__2); +l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___closed__3 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___closed__3(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_reportWF___spec__2___closed__3); l_Lean_Elab_WF_guessLex___lambda__1___closed__1 = _init_l_Lean_Elab_WF_guessLex___lambda__1___closed__1(); lean_mark_persistent(l_Lean_Elab_WF_guessLex___lambda__1___closed__1); l_Lean_Elab_WF_guessLex___lambda__1___closed__2 = _init_l_Lean_Elab_WF_guessLex___lambda__1___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/WF/Main.c b/stage0/stdlib/Lean/Elab/PreDefinition/WF/Main.c index b72cfe1150fc..e2a810736ace 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/WF/Main.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/WF/Main.c @@ -161,7 +161,6 @@ lean_object* l_Lean_Elab_WF_preprocess___lambda__1___boxed(lean_object*, lean_ob lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___spec__5(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_getFixedPrefix___spec__3___lambda__3___closed__4; -lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__6(size_t, size_t, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__10___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2712_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_getFixedPrefix(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -276,6 +275,7 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefi LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Elab_PreDefinition_WF_Main_0__Lean_Elab_addNonRecPreDefs_mkSum___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_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___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___private_Lean_Meta_Basic_0__Lean_Meta_instantiateForallAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__5(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Nat_repr(lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__8___closed__1; @@ -5333,7 +5333,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_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__6(x_2, x_3, x_4); +x_23 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__5(x_2, x_3, x_4); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/WF/TerminationHint.c b/stage0/stdlib/Lean/Elab/PreDefinition/WF/TerminationHint.c index 391832febd26..d1332f3df671 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/WF/TerminationHint.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/WF/TerminationHint.c @@ -14,7 +14,6 @@ extern "C" { #endif lean_object* l___private_Init_Util_0__outOfBounds___rarg(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_elabTerminationHints___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_format_pretty(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_elabTerminationHints___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_TerminationHints_ensureNone___closed__10; @@ -28,16 +27,18 @@ static lean_object* l_Lean_Elab_WF_elabTerminationHints___rarg___lambda__7___clo static lean_object* l_Lean_Elab_WF_elabTerminationHints___rarg___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_WF_elabTerminationHints___rarg___lambda__2(lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_TerminationBy_checkVars___closed__3; +static lean_object* l_Lean_Elab_WF_TerminationBy_unexpand___closed__10; static lean_object* l_Lean_Elab_WF_TerminationBy_checkVars_parameters___closed__1; lean_object* l_Lean_Syntax_formatStxAux(lean_object*, uint8_t, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_WF_TerminationBy_unexpand___closed__16; LEAN_EXPORT lean_object* l_Lean_Elab_WF_TerminationBy_checkVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_toString(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_WF_TerminationBy_checkVars___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getId(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_TerminationHints_rememberExtraParams(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_elabTerminationHints___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* l_Lean_Syntax_getArgs(lean_object*); lean_object* l_Lean_replaceRef(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_TerminationBy_unexpand___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_TerminationHints_ensureNone___closed__7; lean_object* lean_array_fget(lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_TerminationBy_checkVars___closed__7; @@ -54,6 +55,7 @@ static lean_object* l_Lean_Elab_WF_TerminationBy_unexpand___closed__2; lean_object* l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_TerminationBy_checkVars___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_WF_TerminationHints_rememberExtraParams___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_WF_TerminationBy_unexpand___closed__12; static lean_object* l_Lean_Elab_WF_TerminationBy_unexpand___closed__1; static lean_object* l_Lean_Elab_WF_instInhabitedTerminationBy___closed__1; size_t lean_usize_of_nat(lean_object*); @@ -73,11 +75,12 @@ LEAN_EXPORT lean_object* l_Lean_Elab_WF_instInhabitedTerminationBy; lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_TerminationBy_checkVars___closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_WF_TerminationBy_unexpand___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_WF_TerminationBy_unexpand___closed__13; lean_object* l_Lean_Syntax_node3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_WF_TerminationBy_checkVars___closed__12; static lean_object* l_Lean_Elab_WF_TerminationBy_unexpand___closed__5; static lean_object* l_Lean_Elab_WF_elabTerminationHints___rarg___lambda__3___closed__4; static lean_object* l_Lean_Elab_WF_TerminationHints_ensureNone___closed__1; +static lean_object* l_Lean_Elab_WF_TerminationBy_unexpand___closed__15; static lean_object* l_Lean_Elab_WF_elabTerminationHints___rarg___lambda__4___closed__3; static lean_object* l_Lean_Elab_WF_TerminationHints_ensureNone___closed__9; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); @@ -92,7 +95,7 @@ extern lean_object* l_Std_Format_defWidth; static lean_object* l_Lean_Elab_WF_elabTerminationHints___rarg___lambda__6___closed__1; static lean_object* l_Lean_Elab_WF_elabTerminationHints___rarg___lambda__4___closed__6; lean_object* l_Array_append___rarg(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_WF_TerminationBy_checkVars___closed__13; +static lean_object* l_Lean_Elab_WF_TerminationBy_unexpand___closed__11; static lean_object* l_Lean_Elab_WF_TerminationBy_checkVars_parameters___closed__2; static lean_object* l_Lean_Elab_WF_TerminationBy_checkVars_parameters___closed__3; lean_object* l_Lean_throwErrorAt___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -117,6 +120,7 @@ lean_object* l_Lean_logAt___at_Lean_addDecl___spec__6(lean_object*, lean_object* LEAN_EXPORT uint8_t l_Lean_Elab_WF_TerminationBy_synthetic___default; size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Elab_WF_elabTerminationHints___rarg___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_WF_TerminationBy_unexpand___closed__14; lean_object* lean_array_uget(lean_object*, size_t); static lean_object* l_Lean_Elab_WF_elabTerminationHints___rarg___lambda__4___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_WF_elabTerminationHints___rarg___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -133,6 +137,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_WF_TerminationHints_ensureNone___boxed(lean LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_TerminationBy_unexpand___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_TerminationBy_checkVars_parameters(lean_object*); static lean_object* l_Lean_Elab_WF_TerminationHints_ensureNone___closed__2; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_TerminationBy_unexpand___spec__2(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_WF_elabTerminationHints___rarg___lambda__4___closed__2; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); lean_object* l_Lean_MessageData_ofName(lean_object*); @@ -206,6 +211,30 @@ goto _start; } } } +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_TerminationBy_unexpand___spec__2(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; size_t x_9; size_t x_10; lean_object* x_11; +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 = 1; +x_10 = lean_usize_add(x_3, x_9); +x_11 = lean_array_uset(x_8, x_3, x_6); +x_3 = x_10; +x_4 = x_11; +goto _start; +} +} +} static lean_object* _init_l_Lean_Elab_WF_TerminationBy_unexpand___closed__1() { _start: { @@ -280,6 +309,76 @@ static lean_object* _init_l_Lean_Elab_WF_TerminationBy_unexpand___closed__9() { _start: { lean_object* x_1; +x_1 = lean_mk_string_from_bytes("ident", 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_WF_TerminationBy_unexpand___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_WF_TerminationBy_unexpand___closed__9; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_WF_TerminationBy_unexpand___closed__11() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Term", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_WF_TerminationBy_unexpand___closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("hole", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_WF_TerminationBy_unexpand___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_Elab_WF_TerminationBy_unexpand___closed__1; +x_2 = l_Lean_Elab_WF_TerminationBy_unexpand___closed__2; +x_3 = l_Lean_Elab_WF_TerminationBy_unexpand___closed__11; +x_4 = l_Lean_Elab_WF_TerminationBy_unexpand___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_Elab_WF_TerminationBy_unexpand___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_WF_TerminationBy_unexpand___closed__13; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_WF_TerminationBy_unexpand___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_WF_TerminationBy_unexpand___closed__10; +x_2 = l_Lean_Elab_WF_TerminationBy_unexpand___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_Elab_WF_TerminationBy_unexpand___closed__16() { +_start: +{ +lean_object* x_1; x_1 = lean_mk_string_from_bytes("=>", 2); return x_1; } @@ -308,7 +407,7 @@ x_16 = lean_st_ref_get(x_5, x_6); 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_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; size_t x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; x_18 = lean_ctor_get(x_16, 0); lean_dec(x_18); x_19 = l_Lean_Elab_WF_TerminationBy_unexpand___closed__6; @@ -316,129 +415,139 @@ lean_inc(x_15); x_20 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_20, 0, x_15); lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_Elab_WF_instInhabitedTerminationBy___closed__1; -x_22 = l_Array_append___rarg(x_21, x_11); -x_23 = l_Lean_Elab_WF_TerminationBy_unexpand___closed__8; +x_21 = lean_array_get_size(x_11); +x_22 = lean_usize_of_nat(x_21); +lean_dec(x_21); +x_23 = l_Lean_Elab_WF_TerminationBy_unexpand___closed__15; +x_24 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_TerminationBy_unexpand___spec__2(x_23, x_22, x_10, x_11); +x_25 = l_Lean_Elab_WF_instInhabitedTerminationBy___closed__1; +x_26 = l_Array_append___rarg(x_25, x_24); +x_27 = l_Lean_Elab_WF_TerminationBy_unexpand___closed__8; lean_inc(x_15); -x_24 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_24, 0, x_15); -lean_ctor_set(x_24, 1, x_23); -lean_ctor_set(x_24, 2, x_22); -x_25 = l_Lean_Elab_WF_TerminationBy_unexpand___closed__9; +x_28 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_28, 0, x_15); +lean_ctor_set(x_28, 1, x_27); +lean_ctor_set(x_28, 2, x_26); +x_29 = l_Lean_Elab_WF_TerminationBy_unexpand___closed__16; lean_inc(x_15); -x_26 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_26, 0, x_15); -lean_ctor_set(x_26, 1, x_25); +x_30 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_30, 0, x_15); +lean_ctor_set(x_30, 1, x_29); lean_inc(x_15); -x_27 = l_Lean_Syntax_node2(x_15, x_23, x_24, x_26); -x_28 = lean_ctor_get(x_1, 2); -lean_inc(x_28); +x_31 = l_Lean_Syntax_node2(x_15, x_27, x_28, x_30); +x_32 = lean_ctor_get(x_1, 2); +lean_inc(x_32); lean_dec(x_1); -x_29 = l_Lean_Elab_WF_TerminationBy_unexpand___closed__5; -x_30 = l_Lean_Syntax_node3(x_15, x_29, x_20, x_27, x_28); -lean_ctor_set(x_16, 0, x_30); +x_33 = l_Lean_Elab_WF_TerminationBy_unexpand___closed__5; +x_34 = l_Lean_Syntax_node3(x_15, x_33, x_20, x_31, x_32); +lean_ctor_set(x_16, 0, x_34); 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; lean_object* x_37; lean_object* x_38; 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_31 = lean_ctor_get(x_16, 1); -lean_inc(x_31); +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* 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; lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_35 = lean_ctor_get(x_16, 1); +lean_inc(x_35); lean_dec(x_16); -x_32 = l_Lean_Elab_WF_TerminationBy_unexpand___closed__6; -lean_inc(x_15); -x_33 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_33, 0, x_15); -lean_ctor_set(x_33, 1, x_32); -x_34 = l_Lean_Elab_WF_instInhabitedTerminationBy___closed__1; -x_35 = l_Array_append___rarg(x_34, x_11); -x_36 = l_Lean_Elab_WF_TerminationBy_unexpand___closed__8; +x_36 = l_Lean_Elab_WF_TerminationBy_unexpand___closed__6; lean_inc(x_15); -x_37 = lean_alloc_ctor(1, 3, 0); +x_37 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_37, 0, x_15); lean_ctor_set(x_37, 1, x_36); -lean_ctor_set(x_37, 2, x_35); -x_38 = l_Lean_Elab_WF_TerminationBy_unexpand___closed__9; +x_38 = lean_array_get_size(x_11); +x_39 = lean_usize_of_nat(x_38); +lean_dec(x_38); +x_40 = l_Lean_Elab_WF_TerminationBy_unexpand___closed__15; +x_41 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_TerminationBy_unexpand___spec__2(x_40, x_39, x_10, x_11); +x_42 = l_Lean_Elab_WF_instInhabitedTerminationBy___closed__1; +x_43 = l_Array_append___rarg(x_42, x_41); +x_44 = l_Lean_Elab_WF_TerminationBy_unexpand___closed__8; lean_inc(x_15); -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_15); -lean_ctor_set(x_39, 1, x_38); +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_15); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_43); +x_46 = l_Lean_Elab_WF_TerminationBy_unexpand___closed__16; lean_inc(x_15); -x_40 = l_Lean_Syntax_node2(x_15, x_36, x_37, x_39); -x_41 = lean_ctor_get(x_1, 2); -lean_inc(x_41); +x_47 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_47, 0, x_15); +lean_ctor_set(x_47, 1, x_46); +lean_inc(x_15); +x_48 = l_Lean_Syntax_node2(x_15, x_44, x_45, x_47); +x_49 = lean_ctor_get(x_1, 2); +lean_inc(x_49); lean_dec(x_1); -x_42 = l_Lean_Elab_WF_TerminationBy_unexpand___closed__5; -x_43 = l_Lean_Syntax_node3(x_15, x_42, x_33, x_40, x_41); -x_44 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_44, 0, x_43); -lean_ctor_set(x_44, 1, x_31); -return x_44; +x_50 = l_Lean_Elab_WF_TerminationBy_unexpand___closed__5; +x_51 = l_Lean_Syntax_node3(x_15, x_50, x_37, x_48, x_49); +x_52 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_35); +return x_52; } } else { -lean_object* x_45; uint8_t x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; +lean_object* x_53; uint8_t x_54; lean_object* x_55; lean_object* x_56; uint8_t x_57; lean_dec(x_11); -x_45 = lean_ctor_get(x_4, 5); -lean_inc(x_45); +x_53 = lean_ctor_get(x_4, 5); +lean_inc(x_53); lean_dec(x_4); -x_46 = 0; -x_47 = l_Lean_SourceInfo_fromRef(x_45, x_46); -x_48 = lean_st_ref_get(x_5, x_6); -x_49 = !lean_is_exclusive(x_48); -if (x_49 == 0) -{ -lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_50 = lean_ctor_get(x_48, 0); -lean_dec(x_50); -x_51 = l_Lean_Elab_WF_TerminationBy_unexpand___closed__6; -lean_inc(x_47); -x_52 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_52, 0, x_47); -lean_ctor_set(x_52, 1, x_51); -x_53 = l_Lean_Elab_WF_TerminationBy_unexpand___closed__8; -x_54 = l_Lean_Elab_WF_instInhabitedTerminationBy___closed__1; -lean_inc(x_47); -x_55 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_55, 0, x_47); -lean_ctor_set(x_55, 1, x_53); -lean_ctor_set(x_55, 2, x_54); -x_56 = lean_ctor_get(x_1, 2); -lean_inc(x_56); +x_54 = 0; +x_55 = l_Lean_SourceInfo_fromRef(x_53, x_54); +x_56 = lean_st_ref_get(x_5, x_6); +x_57 = !lean_is_exclusive(x_56); +if (x_57 == 0) +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_58 = lean_ctor_get(x_56, 0); +lean_dec(x_58); +x_59 = l_Lean_Elab_WF_TerminationBy_unexpand___closed__6; +lean_inc(x_55); +x_60 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_60, 0, x_55); +lean_ctor_set(x_60, 1, x_59); +x_61 = l_Lean_Elab_WF_TerminationBy_unexpand___closed__8; +x_62 = l_Lean_Elab_WF_instInhabitedTerminationBy___closed__1; +lean_inc(x_55); +x_63 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_63, 0, x_55); +lean_ctor_set(x_63, 1, x_61); +lean_ctor_set(x_63, 2, x_62); +x_64 = lean_ctor_get(x_1, 2); +lean_inc(x_64); lean_dec(x_1); -x_57 = l_Lean_Elab_WF_TerminationBy_unexpand___closed__5; -x_58 = l_Lean_Syntax_node3(x_47, x_57, x_52, x_55, x_56); -lean_ctor_set(x_48, 0, x_58); -return x_48; +x_65 = l_Lean_Elab_WF_TerminationBy_unexpand___closed__5; +x_66 = l_Lean_Syntax_node3(x_55, x_65, x_60, x_63, x_64); +lean_ctor_set(x_56, 0, x_66); +return x_56; } 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; -x_59 = lean_ctor_get(x_48, 1); -lean_inc(x_59); -lean_dec(x_48); -x_60 = l_Lean_Elab_WF_TerminationBy_unexpand___closed__6; -lean_inc(x_47); -x_61 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_61, 0, x_47); -lean_ctor_set(x_61, 1, x_60); -x_62 = l_Lean_Elab_WF_TerminationBy_unexpand___closed__8; -x_63 = l_Lean_Elab_WF_instInhabitedTerminationBy___closed__1; -lean_inc(x_47); -x_64 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_64, 0, x_47); -lean_ctor_set(x_64, 1, x_62); -lean_ctor_set(x_64, 2, x_63); -x_65 = lean_ctor_get(x_1, 2); -lean_inc(x_65); +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_67 = lean_ctor_get(x_56, 1); +lean_inc(x_67); +lean_dec(x_56); +x_68 = l_Lean_Elab_WF_TerminationBy_unexpand___closed__6; +lean_inc(x_55); +x_69 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_69, 0, x_55); +lean_ctor_set(x_69, 1, x_68); +x_70 = l_Lean_Elab_WF_TerminationBy_unexpand___closed__8; +x_71 = l_Lean_Elab_WF_instInhabitedTerminationBy___closed__1; +lean_inc(x_55); +x_72 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_72, 0, x_55); +lean_ctor_set(x_72, 1, x_70); +lean_ctor_set(x_72, 2, x_71); +x_73 = lean_ctor_get(x_1, 2); +lean_inc(x_73); lean_dec(x_1); -x_66 = l_Lean_Elab_WF_TerminationBy_unexpand___closed__5; -x_67 = l_Lean_Syntax_node3(x_47, x_66, x_61, x_64, x_65); -x_68 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_68, 0, x_67); -lean_ctor_set(x_68, 1, x_59); -return x_68; +x_74 = l_Lean_Elab_WF_TerminationBy_unexpand___closed__5; +x_75 = l_Lean_Syntax_node3(x_55, x_74, x_69, x_72, x_73); +x_76 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_76, 0, x_75); +lean_ctor_set(x_76, 1, x_67); +return x_76; } } } @@ -455,6 +564,19 @@ x_6 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_TerminationBy_unexpand___spec__1( return x_6; } } +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_TerminationBy_unexpand___spec__2___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_WF_TerminationBy_unexpand___spec__2(x_1, x_5, x_6, x_4); +lean_dec(x_1); +return x_7; +} +} LEAN_EXPORT lean_object* l_Lean_Elab_WF_TerminationBy_unexpand___boxed(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: { @@ -1071,38 +1193,20 @@ static lean_object* _init_l_Lean_Elab_WF_TerminationBy_checkVars___closed__7() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("ident", 5); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_WF_TerminationBy_checkVars___closed__8() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Elab_WF_TerminationBy_checkVars___closed__7; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Elab_WF_TerminationBy_checkVars___closed__9() { -_start: -{ -lean_object* x_1; x_1 = lean_mk_string_from_bytes(" (Since Lean v4.6.0, the `termination_by` clause no longer ", 59); return x_1; } } -static lean_object* _init_l_Lean_Elab_WF_TerminationBy_checkVars___closed__10() { +static lean_object* _init_l_Lean_Elab_WF_TerminationBy_checkVars___closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_WF_TerminationBy_checkVars___closed__9; +x_1 = l_Lean_Elab_WF_TerminationBy_checkVars___closed__7; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_WF_TerminationBy_checkVars___closed__11() { +static lean_object* _init_l_Lean_Elab_WF_TerminationBy_checkVars___closed__9() { _start: { lean_object* x_1; @@ -1110,21 +1214,21 @@ x_1 = lean_mk_string_from_bytes("expects the function name here.)", 32); return x_1; } } -static lean_object* _init_l_Lean_Elab_WF_TerminationBy_checkVars___closed__12() { +static lean_object* _init_l_Lean_Elab_WF_TerminationBy_checkVars___closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_WF_TerminationBy_checkVars___closed__11; +x_1 = l_Lean_Elab_WF_TerminationBy_checkVars___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_Lean_Elab_WF_TerminationBy_checkVars___closed__13() { +static lean_object* _init_l_Lean_Elab_WF_TerminationBy_checkVars___closed__11() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_WF_TerminationBy_checkVars___closed__12; +x_1 = l_Lean_Elab_WF_TerminationBy_checkVars___closed__10; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; @@ -1202,7 +1306,7 @@ lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; lean_dec(x_10); x_31 = lean_box(0); x_32 = l___private_Init_Util_0__outOfBounds___rarg(x_31); -x_33 = l_Lean_Elab_WF_TerminationBy_checkVars___closed__8; +x_33 = l_Lean_Elab_WF_TerminationBy_unexpand___closed__10; lean_inc(x_32); x_34 = l_Lean_Syntax_isOfKind(x_32, x_33); if (x_34 == 0) @@ -1234,11 +1338,11 @@ return x_40; 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; -x_41 = l_Lean_Elab_WF_TerminationBy_checkVars___closed__10; +x_41 = l_Lean_Elab_WF_TerminationBy_checkVars___closed__8; x_42 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_42, 0, x_28); lean_ctor_set(x_42, 1, x_41); -x_43 = l_Lean_Elab_WF_TerminationBy_checkVars___closed__13; +x_43 = l_Lean_Elab_WF_TerminationBy_checkVars___closed__11; x_44 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_44, 0, x_42); lean_ctor_set(x_44, 1, x_43); @@ -1254,7 +1358,7 @@ else lean_object* x_47; lean_object* x_48; uint8_t x_49; x_47 = lean_array_fget(x_10, x_29); lean_dec(x_10); -x_48 = l_Lean_Elab_WF_TerminationBy_checkVars___closed__8; +x_48 = l_Lean_Elab_WF_TerminationBy_unexpand___closed__10; lean_inc(x_47); x_49 = l_Lean_Syntax_isOfKind(x_47, x_48); if (x_49 == 0) @@ -1286,11 +1390,11 @@ return x_55; else { lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_56 = l_Lean_Elab_WF_TerminationBy_checkVars___closed__10; +x_56 = l_Lean_Elab_WF_TerminationBy_checkVars___closed__8; x_57 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_57, 0, x_28); lean_ctor_set(x_57, 1, x_56); -x_58 = l_Lean_Elab_WF_TerminationBy_checkVars___closed__13; +x_58 = l_Lean_Elab_WF_TerminationBy_checkVars___closed__11; x_59 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_59, 0, x_57); lean_ctor_set(x_59, 1, x_58); @@ -1342,30 +1446,6 @@ lean_dec(x_1); return x_9; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_elabTerminationHints___spec__1(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_lt(x_4, x_3); -if (x_6 == 0) -{ -return x_5; -} -else -{ -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_7 = lean_array_uget(x_5, x_4); -x_8 = lean_unsigned_to_nat(0u); -x_9 = lean_array_uset(x_5, x_4, x_8); -x_10 = 1; -x_11 = lean_usize_add(x_4, x_10); -x_12 = lean_array_uset(x_9, x_4, x_7); -x_4 = x_11; -x_5 = x_12; -goto _start; -} -} -} LEAN_EXPORT lean_object* l_Lean_Elab_WF_elabTerminationHints___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { @@ -1749,198 +1829,184 @@ lean_inc(x_50); x_51 = l_Lean_Syntax_matchesNull(x_50, x_49); 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; size_t x_59; size_t x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; uint8_t x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_dec(x_4); x_52 = l_Lean_Syntax_getArg(x_15, x_32); -x_53 = lean_box(0); -x_54 = l_Lean_Syntax_getArgs(x_50); +x_53 = l_Lean_Syntax_getArgs(x_50); lean_dec(x_50); -x_55 = lean_ctor_get(x_1, 1); +x_54 = lean_ctor_get(x_1, 1); +lean_inc(x_54); +x_55 = lean_ctor_get(x_1, 0); lean_inc(x_55); -x_56 = lean_ctor_get(x_1, 0); -lean_inc(x_56); lean_dec(x_1); -x_57 = lean_ctor_get(x_56, 1); -lean_inc(x_57); -lean_dec(x_56); -x_58 = lean_array_get_size(x_54); -x_59 = lean_usize_of_nat(x_58); -lean_dec(x_58); -x_60 = 0; -x_61 = l_Lean_Elab_WF_TerminationBy_checkVars___closed__8; -x_62 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_elabTerminationHints___spec__1(x_53, x_61, x_59, x_60, x_54); -x_63 = 0; -x_64 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_64, 0, x_15); -lean_ctor_set(x_64, 1, x_62); -lean_ctor_set(x_64, 2, x_52); -lean_ctor_set_uint8(x_64, sizeof(void*)*3, x_63); -lean_ctor_set(x_5, 0, x_64); -x_65 = lean_apply_2(x_57, lean_box(0), x_5); -x_66 = lean_apply_4(x_55, lean_box(0), lean_box(0), x_65, x_7); -return x_66; +x_56 = lean_ctor_get(x_55, 1); +lean_inc(x_56); +lean_dec(x_55); +x_57 = 0; +x_58 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_58, 0, x_15); +lean_ctor_set(x_58, 1, x_53); +lean_ctor_set(x_58, 2, x_52); +lean_ctor_set_uint8(x_58, sizeof(void*)*3, x_57); +lean_ctor_set(x_5, 0, x_58); +x_59 = lean_apply_2(x_56, lean_box(0), x_5); +x_60 = lean_apply_4(x_54, lean_box(0), lean_box(0), x_59, x_7); +return x_60; } else { -lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_dec(x_50); lean_free_object(x_5); -x_67 = lean_ctor_get(x_1, 1); -lean_inc(x_67); -x_68 = l_Lean_Elab_WF_elabTerminationHints___rarg___lambda__4___closed__6; -x_69 = l_Lean_throwErrorAt___rarg(x_1, x_4, x_15, x_68); -x_70 = lean_apply_4(x_67, lean_box(0), lean_box(0), x_69, x_7); -return x_70; +x_61 = lean_ctor_get(x_1, 1); +lean_inc(x_61); +x_62 = l_Lean_Elab_WF_elabTerminationHints___rarg___lambda__4___closed__6; +x_63 = l_Lean_throwErrorAt___rarg(x_1, x_4, x_15, x_62); +x_64 = lean_apply_4(x_61, lean_box(0), lean_box(0), x_63, x_7); +return x_64; } } } } else { -lean_object* x_71; lean_object* x_72; uint8_t x_73; -x_71 = lean_ctor_get(x_5, 0); -lean_inc(x_71); +lean_object* x_65; lean_object* x_66; uint8_t x_67; +x_65 = lean_ctor_get(x_5, 0); +lean_inc(x_65); lean_dec(x_5); -x_72 = l_Lean_Elab_WF_TerminationBy_unexpand___closed__5; -lean_inc(x_71); -x_73 = l_Lean_Syntax_isOfKind(x_71, x_72); -if (x_73 == 0) -{ -lean_object* x_74; uint8_t x_75; -x_74 = l_Lean_Elab_WF_elabTerminationHints___rarg___lambda__4___closed__2; -lean_inc(x_71); -x_75 = l_Lean_Syntax_isOfKind(x_71, x_74); -if (x_75 == 0) -{ -lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; -x_76 = lean_ctor_get(x_1, 1); -lean_inc(x_76); -x_77 = l_Lean_Elab_WF_elabTerminationHints___rarg___lambda__4___closed__4; -x_78 = l_Lean_throwErrorAt___rarg(x_1, x_4, x_71, x_77); -x_79 = lean_apply_4(x_76, lean_box(0), lean_box(0), x_78, x_7); -return x_79; +x_66 = l_Lean_Elab_WF_TerminationBy_unexpand___closed__5; +lean_inc(x_65); +x_67 = l_Lean_Syntax_isOfKind(x_65, x_66); +if (x_67 == 0) +{ +lean_object* x_68; uint8_t x_69; +x_68 = l_Lean_Elab_WF_elabTerminationHints___rarg___lambda__4___closed__2; +lean_inc(x_65); +x_69 = l_Lean_Syntax_isOfKind(x_65, x_68); +if (x_69 == 0) +{ +lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_70 = lean_ctor_get(x_1, 1); +lean_inc(x_70); +x_71 = l_Lean_Elab_WF_elabTerminationHints___rarg___lambda__4___closed__4; +x_72 = l_Lean_throwErrorAt___rarg(x_1, x_4, x_65, x_71); +x_73 = lean_apply_4(x_70, lean_box(0), lean_box(0), x_72, x_7); +return x_73; } 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_dec(x_71); +lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +lean_dec(x_65); lean_dec(x_4); -x_80 = lean_ctor_get(x_1, 1); -lean_inc(x_80); -x_81 = lean_ctor_get(x_1, 0); -lean_inc(x_81); +x_74 = lean_ctor_get(x_1, 1); +lean_inc(x_74); +x_75 = lean_ctor_get(x_1, 0); +lean_inc(x_75); lean_dec(x_1); -x_82 = lean_ctor_get(x_81, 1); -lean_inc(x_82); -lean_dec(x_81); -x_83 = lean_box(0); -x_84 = lean_apply_2(x_82, lean_box(0), x_83); -x_85 = lean_apply_4(x_80, lean_box(0), lean_box(0), x_84, x_7); -return x_85; +x_76 = lean_ctor_get(x_75, 1); +lean_inc(x_76); +lean_dec(x_75); +x_77 = lean_box(0); +x_78 = lean_apply_2(x_76, lean_box(0), x_77); +x_79 = lean_apply_4(x_74, lean_box(0), lean_box(0), x_78, x_7); +return x_79; } } else { -lean_object* x_86; lean_object* x_87; lean_object* x_88; uint8_t x_89; -x_86 = lean_unsigned_to_nat(1u); -x_87 = l_Lean_Syntax_getArg(x_71, x_86); -x_88 = lean_unsigned_to_nat(2u); -lean_inc(x_87); -x_89 = l_Lean_Syntax_matchesNull(x_87, x_88); -if (x_89 == 0) -{ -lean_object* x_90; uint8_t x_91; -x_90 = lean_unsigned_to_nat(0u); -x_91 = l_Lean_Syntax_matchesNull(x_87, x_90); -if (x_91 == 0) -{ -lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; -x_92 = lean_ctor_get(x_1, 1); -lean_inc(x_92); -x_93 = l_Lean_Elab_WF_elabTerminationHints___rarg___lambda__4___closed__4; -x_94 = l_Lean_throwErrorAt___rarg(x_1, x_4, x_71, x_93); -x_95 = lean_apply_4(x_92, lean_box(0), lean_box(0), x_94, x_7); -return x_95; +lean_object* x_80; lean_object* x_81; lean_object* x_82; uint8_t x_83; +x_80 = lean_unsigned_to_nat(1u); +x_81 = l_Lean_Syntax_getArg(x_65, x_80); +x_82 = lean_unsigned_to_nat(2u); +lean_inc(x_81); +x_83 = l_Lean_Syntax_matchesNull(x_81, x_82); +if (x_83 == 0) +{ +lean_object* x_84; uint8_t x_85; +x_84 = lean_unsigned_to_nat(0u); +x_85 = l_Lean_Syntax_matchesNull(x_81, x_84); +if (x_85 == 0) +{ +lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_86 = lean_ctor_get(x_1, 1); +lean_inc(x_86); +x_87 = l_Lean_Elab_WF_elabTerminationHints___rarg___lambda__4___closed__4; +x_88 = l_Lean_throwErrorAt___rarg(x_1, x_4, x_65, x_87); +x_89 = lean_apply_4(x_86, lean_box(0), lean_box(0), x_88, x_7); +return x_89; } else { -lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; uint8_t x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; +lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; uint8_t x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_dec(x_4); -x_96 = l_Lean_Syntax_getArg(x_71, x_88); -x_97 = lean_ctor_get(x_1, 1); -lean_inc(x_97); -x_98 = lean_ctor_get(x_1, 0); -lean_inc(x_98); +x_90 = l_Lean_Syntax_getArg(x_65, x_82); +x_91 = lean_ctor_get(x_1, 1); +lean_inc(x_91); +x_92 = lean_ctor_get(x_1, 0); +lean_inc(x_92); lean_dec(x_1); -x_99 = lean_ctor_get(x_98, 1); -lean_inc(x_99); -lean_dec(x_98); -x_100 = l_Lean_Elab_WF_instInhabitedTerminationBy___closed__1; -x_101 = 0; -x_102 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_102, 0, x_71); -lean_ctor_set(x_102, 1, x_100); -lean_ctor_set(x_102, 2, x_96); -lean_ctor_set_uint8(x_102, sizeof(void*)*3, x_101); -x_103 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_103, 0, x_102); -x_104 = lean_apply_2(x_99, lean_box(0), x_103); -x_105 = lean_apply_4(x_97, lean_box(0), lean_box(0), x_104, x_7); -return x_105; +x_93 = lean_ctor_get(x_92, 1); +lean_inc(x_93); +lean_dec(x_92); +x_94 = l_Lean_Elab_WF_instInhabitedTerminationBy___closed__1; +x_95 = 0; +x_96 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_96, 0, x_65); +lean_ctor_set(x_96, 1, x_94); +lean_ctor_set(x_96, 2, x_90); +lean_ctor_set_uint8(x_96, sizeof(void*)*3, x_95); +x_97 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_97, 0, x_96); +x_98 = lean_apply_2(x_93, lean_box(0), x_97); +x_99 = lean_apply_4(x_91, lean_box(0), lean_box(0), x_98, x_7); +return x_99; } } else { -lean_object* x_106; lean_object* x_107; uint8_t x_108; -x_106 = lean_unsigned_to_nat(0u); -x_107 = l_Lean_Syntax_getArg(x_87, x_106); -lean_dec(x_87); -lean_inc(x_107); -x_108 = l_Lean_Syntax_matchesNull(x_107, x_106); -if (x_108 == 0) +lean_object* x_100; lean_object* x_101; uint8_t x_102; +x_100 = lean_unsigned_to_nat(0u); +x_101 = l_Lean_Syntax_getArg(x_81, x_100); +lean_dec(x_81); +lean_inc(x_101); +x_102 = l_Lean_Syntax_matchesNull(x_101, x_100); +if (x_102 == 0) { -lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; size_t x_116; size_t x_117; lean_object* x_118; lean_object* x_119; uint8_t x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; +lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; uint8_t x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_dec(x_4); -x_109 = l_Lean_Syntax_getArg(x_71, x_88); -x_110 = lean_box(0); -x_111 = l_Lean_Syntax_getArgs(x_107); -lean_dec(x_107); -x_112 = lean_ctor_get(x_1, 1); -lean_inc(x_112); -x_113 = lean_ctor_get(x_1, 0); -lean_inc(x_113); +x_103 = l_Lean_Syntax_getArg(x_65, x_82); +x_104 = l_Lean_Syntax_getArgs(x_101); +lean_dec(x_101); +x_105 = lean_ctor_get(x_1, 1); +lean_inc(x_105); +x_106 = lean_ctor_get(x_1, 0); +lean_inc(x_106); lean_dec(x_1); -x_114 = lean_ctor_get(x_113, 1); -lean_inc(x_114); -lean_dec(x_113); -x_115 = lean_array_get_size(x_111); -x_116 = lean_usize_of_nat(x_115); -lean_dec(x_115); -x_117 = 0; -x_118 = l_Lean_Elab_WF_TerminationBy_checkVars___closed__8; -x_119 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_elabTerminationHints___spec__1(x_110, x_118, x_116, x_117, x_111); -x_120 = 0; -x_121 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_121, 0, x_71); -lean_ctor_set(x_121, 1, x_119); -lean_ctor_set(x_121, 2, x_109); -lean_ctor_set_uint8(x_121, sizeof(void*)*3, x_120); -x_122 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_122, 0, x_121); -x_123 = lean_apply_2(x_114, lean_box(0), x_122); -x_124 = lean_apply_4(x_112, lean_box(0), lean_box(0), x_123, x_7); -return x_124; +x_107 = lean_ctor_get(x_106, 1); +lean_inc(x_107); +lean_dec(x_106); +x_108 = 0; +x_109 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_109, 0, x_65); +lean_ctor_set(x_109, 1, x_104); +lean_ctor_set(x_109, 2, x_103); +lean_ctor_set_uint8(x_109, sizeof(void*)*3, x_108); +x_110 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_110, 0, x_109); +x_111 = lean_apply_2(x_107, lean_box(0), x_110); +x_112 = lean_apply_4(x_105, lean_box(0), lean_box(0), x_111, x_7); +return x_112; } else { -lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; -lean_dec(x_107); -x_125 = lean_ctor_get(x_1, 1); -lean_inc(x_125); -x_126 = l_Lean_Elab_WF_elabTerminationHints___rarg___lambda__4___closed__6; -x_127 = l_Lean_throwErrorAt___rarg(x_1, x_4, x_71, x_126); -x_128 = lean_apply_4(x_125, lean_box(0), lean_box(0), x_127, x_7); -return x_128; +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +lean_dec(x_101); +x_113 = lean_ctor_get(x_1, 1); +lean_inc(x_113); +x_114 = l_Lean_Elab_WF_elabTerminationHints___rarg___lambda__4___closed__6; +x_115 = l_Lean_throwErrorAt___rarg(x_1, x_4, x_65, x_114); +x_116 = lean_apply_4(x_113, lean_box(0), lean_box(0), x_115, x_7); +return x_116; } } } @@ -2373,20 +2439,6 @@ x_2 = lean_alloc_closure((void*)(l_Lean_Elab_WF_elabTerminationHints___rarg), 3, return x_2; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_elabTerminationHints___spec__1___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_mapMUnsafe_map___at_Lean_Elab_WF_elabTerminationHints___spec__1(x_1, x_2, x_6, x_7, x_5); -lean_dec(x_2); -lean_dec(x_1); -return x_8; -} -} LEAN_EXPORT lean_object* l_Lean_Elab_WF_elabTerminationHints___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) { _start: { @@ -2430,6 +2482,20 @@ l_Lean_Elab_WF_TerminationBy_unexpand___closed__8 = _init_l_Lean_Elab_WF_Termina lean_mark_persistent(l_Lean_Elab_WF_TerminationBy_unexpand___closed__8); l_Lean_Elab_WF_TerminationBy_unexpand___closed__9 = _init_l_Lean_Elab_WF_TerminationBy_unexpand___closed__9(); lean_mark_persistent(l_Lean_Elab_WF_TerminationBy_unexpand___closed__9); +l_Lean_Elab_WF_TerminationBy_unexpand___closed__10 = _init_l_Lean_Elab_WF_TerminationBy_unexpand___closed__10(); +lean_mark_persistent(l_Lean_Elab_WF_TerminationBy_unexpand___closed__10); +l_Lean_Elab_WF_TerminationBy_unexpand___closed__11 = _init_l_Lean_Elab_WF_TerminationBy_unexpand___closed__11(); +lean_mark_persistent(l_Lean_Elab_WF_TerminationBy_unexpand___closed__11); +l_Lean_Elab_WF_TerminationBy_unexpand___closed__12 = _init_l_Lean_Elab_WF_TerminationBy_unexpand___closed__12(); +lean_mark_persistent(l_Lean_Elab_WF_TerminationBy_unexpand___closed__12); +l_Lean_Elab_WF_TerminationBy_unexpand___closed__13 = _init_l_Lean_Elab_WF_TerminationBy_unexpand___closed__13(); +lean_mark_persistent(l_Lean_Elab_WF_TerminationBy_unexpand___closed__13); +l_Lean_Elab_WF_TerminationBy_unexpand___closed__14 = _init_l_Lean_Elab_WF_TerminationBy_unexpand___closed__14(); +lean_mark_persistent(l_Lean_Elab_WF_TerminationBy_unexpand___closed__14); +l_Lean_Elab_WF_TerminationBy_unexpand___closed__15 = _init_l_Lean_Elab_WF_TerminationBy_unexpand___closed__15(); +lean_mark_persistent(l_Lean_Elab_WF_TerminationBy_unexpand___closed__15); +l_Lean_Elab_WF_TerminationBy_unexpand___closed__16 = _init_l_Lean_Elab_WF_TerminationBy_unexpand___closed__16(); +lean_mark_persistent(l_Lean_Elab_WF_TerminationBy_unexpand___closed__16); l_Lean_Elab_WF_instInhabitedDecreasingBy___closed__1 = _init_l_Lean_Elab_WF_instInhabitedDecreasingBy___closed__1(); lean_mark_persistent(l_Lean_Elab_WF_instInhabitedDecreasingBy___closed__1); l_Lean_Elab_WF_instInhabitedDecreasingBy = _init_l_Lean_Elab_WF_instInhabitedDecreasingBy(); @@ -2492,10 +2558,6 @@ l_Lean_Elab_WF_TerminationBy_checkVars___closed__10 = _init_l_Lean_Elab_WF_Termi lean_mark_persistent(l_Lean_Elab_WF_TerminationBy_checkVars___closed__10); l_Lean_Elab_WF_TerminationBy_checkVars___closed__11 = _init_l_Lean_Elab_WF_TerminationBy_checkVars___closed__11(); lean_mark_persistent(l_Lean_Elab_WF_TerminationBy_checkVars___closed__11); -l_Lean_Elab_WF_TerminationBy_checkVars___closed__12 = _init_l_Lean_Elab_WF_TerminationBy_checkVars___closed__12(); -lean_mark_persistent(l_Lean_Elab_WF_TerminationBy_checkVars___closed__12); -l_Lean_Elab_WF_TerminationBy_checkVars___closed__13 = _init_l_Lean_Elab_WF_TerminationBy_checkVars___closed__13(); -lean_mark_persistent(l_Lean_Elab_WF_TerminationBy_checkVars___closed__13); l_Lean_Elab_WF_elabTerminationHints___rarg___lambda__3___closed__1 = _init_l_Lean_Elab_WF_elabTerminationHints___rarg___lambda__3___closed__1(); lean_mark_persistent(l_Lean_Elab_WF_elabTerminationHints___rarg___lambda__3___closed__1); l_Lean_Elab_WF_elabTerminationHints___rarg___lambda__3___closed__2 = _init_l_Lean_Elab_WF_elabTerminationHints___rarg___lambda__3___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/Syntax.c b/stage0/stdlib/Lean/Elab/Syntax.c index b203defaee63..f68494a6a399 100644 --- a/stage0/stdlib/Lean/Elab/Syntax.c +++ b/stage0/stdlib/Lean/Elab/Syntax.c @@ -53,6 +53,7 @@ static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_toParserDescr_pro static lean_object* l_panic___at_Lean_Elab_Term_toParserDescr_processAlias___spec__2___closed__2; LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabSyntax___spec__11(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processAlias(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__12(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_withNotFirst(lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__39; static lean_object* l_Lean_Elab_Term_toParserDescr_process___closed__6; @@ -542,7 +543,6 @@ 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; @@ -15036,7 +15036,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_runLinters___spec__13(x_6, x_3, x_4, x_5); +x_7 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__12(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); diff --git a/stage0/stdlib/Lean/Elab/Tactic/Basic.c b/stage0/stdlib/Lean/Elab/Tactic/Basic.c index 60664b8d764d..e72e43aaf039 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Basic.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Basic.c @@ -13,7 +13,6 @@ #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; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getMainModule(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -26,7 +25,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_initFn____x40_Lean_Elab_Tactic_Basic___hyg_5459____closed__15; static lean_object* l_Lean_Elab_Tactic_instAlternativeTacticM___lambda__1___closed__1; @@ -39,12 +37,12 @@ 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*); static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_5459____closed__3; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__4(lean_object*, lean_object*, uint8_t, lean_object*, 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_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; +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Tactic_evalTactic___spec__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_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*); static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__5___closed__1; @@ -72,6 +70,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_liftMetaTactic___lambda__1(lean_obje lean_object* l_Lean_FileMap_toPosition(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withInfoTreeContext___at_Lean_Elab_Tactic_withTacticInfoContext___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_MessageData_joinSep(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___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*); static lean_object* l_Lean_Elab_throwAbortTactic___at_Lean_Elab_Tactic_done___spec__1___rarg___closed__1; 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*); @@ -98,6 +97,7 @@ LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_evalTa LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_mkTacticInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Elab_Term_instMonadTermElabM; lean_object* l_Lean_Syntax_getArgs(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___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*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__5___boxed(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*); @@ -215,6 +215,7 @@ 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_5459____closed__14; static lean_object* l_Lean_Elab_Tactic_evalTactic___lambda__4___closed__1; +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Tactic_evalTactic___spec__10___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*); 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*); @@ -260,7 +261,6 @@ 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; @@ -298,6 +298,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Basic_0__Lean_Elab_Tactic_ 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_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*); +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Tactic_evalTactic___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_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*); uint8_t l_Lean_Elab_isAbortTacticException(lean_object*); @@ -314,6 +315,7 @@ 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_object* l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNode___spec__1(size_t, size_t, 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*); @@ -325,7 +327,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_closeUsingOrAdmit(lean_object*, lean 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*); @@ -339,11 +340,11 @@ static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hy 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; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_evalTactic___spec__8(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_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*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTacticAtRaw(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_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*); @@ -367,7 +368,6 @@ 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*); @@ -375,10 +375,8 @@ LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_evalTa 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; -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*); @@ -424,7 +422,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_withCaseRef___rarg___lambda__1___box lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_liftMetaFinishingTactic(lean_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__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*); @@ -443,9 +440,11 @@ 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*); +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Tactic_evalTactic___spec__10(lean_object*); uint8_t l___private_Lean_Message_0__Lean_beqMessageSeverity____x40_Lean_Message___hyg_103_(uint8_t, uint8_t); 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*); +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_evalTactic___spec__9(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_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; @@ -456,6 +455,7 @@ 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_MonadExcept_ofExcept___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_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_evalTactic___lambda__4___closed__6; uint8_t l_Lean_Exception_isRuntime(lean_object*); @@ -472,13 +472,11 @@ 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*); @@ -6824,125 +6822,6 @@ 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: { @@ -6950,76 +6829,269 @@ 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; +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; size_t x_24; size_t x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; 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_18 = lean_st_ref_get(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); +x_21 = lean_ctor_get(x_19, 3); +lean_inc(x_21); +lean_dec(x_19); +x_22 = l_Lean_PersistentArray_toArray___rarg(x_21); +x_23 = lean_array_get_size(x_22); +x_24 = lean_usize_of_nat(x_23); +lean_dec(x_23); +x_25 = 0; +x_26 = l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNode___spec__1(x_24, x_25, x_22); +x_27 = lean_alloc_ctor(9, 3, 1); +lean_ctor_set(x_27, 0, x_2); +lean_ctor_set(x_27, 1, x_4); +lean_ctor_set(x_27, 2, x_26); +lean_ctor_set_uint8(x_27, sizeof(void*)*3, x_5); +x_28 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_27, x_10, x_11, x_12, x_13, x_20); lean_dec(x_12); -return x_21; +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); +x_31 = lean_st_ref_take(x_13, x_30); +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 = !lean_is_exclusive(x_32); +if (x_34 == 0) +{ +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_32, 3); +lean_dec(x_35); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_3); +lean_ctor_set(x_36, 1, x_29); +x_37 = l_Lean_PersistentArray_push___rarg(x_1, x_36); +lean_ctor_set(x_32, 3, x_37); +x_38 = lean_st_ref_set(x_13, x_32, x_33); +x_39 = !lean_is_exclusive(x_38); +if (x_39 == 0) +{ +lean_object* x_40; lean_object* x_41; +x_40 = lean_ctor_get(x_38, 0); +lean_dec(x_40); +x_41 = lean_box(0); +lean_ctor_set(x_38, 0, x_41); +return x_38; } 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; uint8_t x_33; 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_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); -x_33 = lean_ctor_get_uint8(x_12, sizeof(void*)*11); -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_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_38, 1); +lean_inc(x_42); +lean_dec(x_38); +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_42); +return x_44; +} +} +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; 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_45 = lean_ctor_get(x_32, 0); +x_46 = lean_ctor_get(x_32, 1); +x_47 = lean_ctor_get(x_32, 2); +x_48 = lean_ctor_get(x_32, 4); +x_49 = lean_ctor_get(x_32, 5); +x_50 = lean_ctor_get(x_32, 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_dec(x_32); +x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_51, 0, x_3); +lean_ctor_set(x_51, 1, x_29); +x_52 = l_Lean_PersistentArray_push___rarg(x_1, x_51); +x_53 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_53, 0, x_45); +lean_ctor_set(x_53, 1, x_46); +lean_ctor_set(x_53, 2, x_47); +lean_ctor_set(x_53, 3, x_52); +lean_ctor_set(x_53, 4, x_48); +lean_ctor_set(x_53, 5, x_49); +lean_ctor_set(x_53, 6, x_50); +x_54 = lean_st_ref_set(x_13, x_53, x_33); +x_55 = lean_ctor_get(x_54, 1); +lean_inc(x_55); +if (lean_is_exclusive(x_54)) { + lean_ctor_release(x_54, 0); + lean_ctor_release(x_54, 1); + x_56 = x_54; +} else { + lean_dec_ref(x_54); + x_56 = lean_box(0); +} +x_57 = lean_box(0); +if (lean_is_scalar(x_56)) { + x_58 = lean_alloc_ctor(0, 2, 0); +} else { + x_58 = x_56; +} +lean_ctor_set(x_58, 0, x_57); +lean_ctor_set(x_58, 1, x_55); +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; uint8_t x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; size_t x_79; size_t x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; +x_59 = lean_ctor_get(x_12, 0); +x_60 = lean_ctor_get(x_12, 1); +x_61 = lean_ctor_get(x_12, 2); +x_62 = lean_ctor_get(x_12, 3); +x_63 = lean_ctor_get(x_12, 4); +x_64 = lean_ctor_get(x_12, 5); +x_65 = lean_ctor_get(x_12, 6); +x_66 = lean_ctor_get(x_12, 7); +x_67 = lean_ctor_get(x_12, 8); +x_68 = lean_ctor_get(x_12, 9); +x_69 = lean_ctor_get(x_12, 10); +x_70 = lean_ctor_get_uint8(x_12, sizeof(void*)*11); +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_dec(x_12); -x_34 = l_Lean_replaceRef(x_3, x_27); -lean_dec(x_27); -x_35 = lean_alloc_ctor(0, 11, 1); -lean_ctor_set(x_35, 0, x_22); -lean_ctor_set(x_35, 1, x_23); -lean_ctor_set(x_35, 2, x_24); -lean_ctor_set(x_35, 3, x_25); -lean_ctor_set(x_35, 4, x_26); -lean_ctor_set(x_35, 5, x_34); -lean_ctor_set(x_35, 6, x_28); -lean_ctor_set(x_35, 7, x_29); -lean_ctor_set(x_35, 8, x_30); -lean_ctor_set(x_35, 9, x_31); -lean_ctor_set(x_35, 10, x_32); -lean_ctor_set_uint8(x_35, sizeof(void*)*11, x_33); -x_36 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_4, x_10, x_11, x_35, x_13, x_14); -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___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Elab_Tactic_evalTactic___spec__7(x_1, x_2, x_3, x_37, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_35, x_13, x_38); -lean_dec(x_35); -return x_39; +x_71 = l_Lean_replaceRef(x_3, x_64); +lean_dec(x_64); +x_72 = lean_alloc_ctor(0, 11, 1); +lean_ctor_set(x_72, 0, x_59); +lean_ctor_set(x_72, 1, x_60); +lean_ctor_set(x_72, 2, x_61); +lean_ctor_set(x_72, 3, x_62); +lean_ctor_set(x_72, 4, x_63); +lean_ctor_set(x_72, 5, x_71); +lean_ctor_set(x_72, 6, x_65); +lean_ctor_set(x_72, 7, x_66); +lean_ctor_set(x_72, 8, x_67); +lean_ctor_set(x_72, 9, x_68); +lean_ctor_set(x_72, 10, x_69); +lean_ctor_set_uint8(x_72, sizeof(void*)*11, x_70); +x_73 = lean_st_ref_get(x_13, x_14); +x_74 = lean_ctor_get(x_73, 0); +lean_inc(x_74); +x_75 = lean_ctor_get(x_73, 1); +lean_inc(x_75); +lean_dec(x_73); +x_76 = lean_ctor_get(x_74, 3); +lean_inc(x_76); +lean_dec(x_74); +x_77 = l_Lean_PersistentArray_toArray___rarg(x_76); +x_78 = lean_array_get_size(x_77); +x_79 = lean_usize_of_nat(x_78); +lean_dec(x_78); +x_80 = 0; +x_81 = l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNode___spec__1(x_79, x_80, x_77); +x_82 = lean_alloc_ctor(9, 3, 1); +lean_ctor_set(x_82, 0, x_2); +lean_ctor_set(x_82, 1, x_4); +lean_ctor_set(x_82, 2, x_81); +lean_ctor_set_uint8(x_82, sizeof(void*)*3, x_5); +x_83 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_82, x_10, x_11, x_72, x_13, x_75); +lean_dec(x_72); +x_84 = lean_ctor_get(x_83, 0); +lean_inc(x_84); +x_85 = lean_ctor_get(x_83, 1); +lean_inc(x_85); +lean_dec(x_83); +x_86 = lean_st_ref_take(x_13, x_85); +x_87 = lean_ctor_get(x_86, 0); +lean_inc(x_87); +x_88 = lean_ctor_get(x_86, 1); +lean_inc(x_88); +lean_dec(x_86); +x_89 = lean_ctor_get(x_87, 0); +lean_inc(x_89); +x_90 = lean_ctor_get(x_87, 1); +lean_inc(x_90); +x_91 = lean_ctor_get(x_87, 2); +lean_inc(x_91); +x_92 = lean_ctor_get(x_87, 4); +lean_inc(x_92); +x_93 = lean_ctor_get(x_87, 5); +lean_inc(x_93); +x_94 = lean_ctor_get(x_87, 6); +lean_inc(x_94); +if (lean_is_exclusive(x_87)) { + lean_ctor_release(x_87, 0); + lean_ctor_release(x_87, 1); + lean_ctor_release(x_87, 2); + lean_ctor_release(x_87, 3); + lean_ctor_release(x_87, 4); + lean_ctor_release(x_87, 5); + lean_ctor_release(x_87, 6); + x_95 = x_87; +} else { + lean_dec_ref(x_87); + x_95 = lean_box(0); +} +x_96 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_96, 0, x_3); +lean_ctor_set(x_96, 1, x_84); +x_97 = l_Lean_PersistentArray_push___rarg(x_1, x_96); +if (lean_is_scalar(x_95)) { + x_98 = lean_alloc_ctor(0, 7, 0); +} else { + x_98 = x_95; +} +lean_ctor_set(x_98, 0, x_89); +lean_ctor_set(x_98, 1, x_90); +lean_ctor_set(x_98, 2, x_91); +lean_ctor_set(x_98, 3, x_97); +lean_ctor_set(x_98, 4, x_92); +lean_ctor_set(x_98, 5, x_93); +lean_ctor_set(x_98, 6, x_94); +x_99 = lean_st_ref_set(x_13, x_98, x_88); +x_100 = lean_ctor_get(x_99, 1); +lean_inc(x_100); +if (lean_is_exclusive(x_99)) { + lean_ctor_release(x_99, 0); + lean_ctor_release(x_99, 1); + x_101 = x_99; +} else { + lean_dec_ref(x_99); + x_101 = lean_box(0); } +x_102 = lean_box(0); +if (lean_is_scalar(x_101)) { + x_103 = lean_alloc_ctor(0, 2, 0); +} else { + x_103 = x_101; } +lean_ctor_set(x_103, 0, x_102); +lean_ctor_set(x_103, 1, x_100); +return x_103; } -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) { +} +} +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Tactic_evalTactic___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { if (lean_obj_tag(x_1) == 0) @@ -7119,7 +7191,7 @@ x_17 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Tactic_ev 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); +x_19 = l_MonadExcept_ofExcept___at_Lean_Elab_Tactic_evalTactic___spec__7(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); @@ -7401,7 +7473,7 @@ x_34 = lean_st_ref_set(x_16, x_29, x_30); x_35 = lean_ctor_get(x_34, 1); lean_inc(x_35); lean_dec(x_34); -x_36 = l_MonadExcept_ofExcept___at_Lean_Elab_Tactic_evalTactic___spec__8(x_25, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_35); +x_36 = l_MonadExcept_ofExcept___at_Lean_Elab_Tactic_evalTactic___spec__7(x_25, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_35); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); @@ -7487,7 +7559,7 @@ x_54 = lean_st_ref_set(x_16, x_53, x_30); x_55 = lean_ctor_get(x_54, 1); lean_inc(x_55); lean_dec(x_54); -x_56 = l_MonadExcept_ofExcept___at_Lean_Elab_Tactic_evalTactic___spec__8(x_25, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_55); +x_56 = l_MonadExcept_ofExcept___at_Lean_Elab_Tactic_evalTactic___spec__7(x_25, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_55); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); @@ -7700,7 +7772,7 @@ 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) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_evalTactic___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, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { uint8_t x_14; @@ -7783,7 +7855,7 @@ 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) { +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_evalTactic___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; lean_object* x_12; lean_object* x_13; @@ -7797,7 +7869,7 @@ 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) { +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Tactic_evalTactic___spec__10___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; @@ -7806,11 +7878,11 @@ 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) { +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Tactic_evalTactic___spec__10(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); +x_2 = lean_alloc_closure((void*)(l_Lean_profileitM___at_Lean_Elab_Tactic_evalTactic___spec__10___rarg___boxed), 13, 0); return x_2; } } @@ -8212,7 +8284,7 @@ lean_dec(x_42); x_51 = lean_box(0); x_52 = l_Lean_Elab_Tactic_evalTactic___lambda__5___boxed__const__1; x_53 = lean_box_usize(x_50); -x_54 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_evalTactic___spec__9___boxed), 13, 10); +x_54 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_evalTactic___spec__8___boxed), 13, 10); lean_closure_set(x_54, 0, x_41); lean_closure_set(x_54, 1, x_52); lean_closure_set(x_54, 2, x_53); @@ -8269,7 +8341,7 @@ lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_1); -x_64 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_evalTactic___spec__10(x_24, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_64 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_evalTactic___spec__9(x_24, 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); @@ -8452,7 +8524,7 @@ lean_dec(x_96); x_105 = lean_box(0); x_106 = l_Lean_Elab_Tactic_evalTactic___lambda__5___boxed__const__1; x_107 = lean_box_usize(x_104); -x_108 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_evalTactic___spec__9___boxed), 13, 10); +x_108 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_evalTactic___spec__8___boxed), 13, 10); lean_closure_set(x_108, 0, x_95); lean_closure_set(x_108, 1, x_106); lean_closure_set(x_108, 2, x_107); @@ -8509,7 +8581,7 @@ lean_dec(x_67); lean_dec(x_66); lean_dec(x_65); lean_dec(x_1); -x_118 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_evalTactic___spec__10(x_77, x_2, x_3, x_4, x_5, x_6, x_7, x_78, x_9, x_10); +x_118 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_evalTactic___spec__9(x_77, x_2, x_3, x_4, x_5, x_6, x_7, x_78, x_9, x_10); lean_dec(x_9); lean_dec(x_78); lean_dec(x_7); @@ -8542,7 +8614,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__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); +x_15 = l_Lean_profileitM___at_Lean_Elab_Tactic_evalTactic___spec__10___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; } @@ -8687,24 +8759,6 @@ 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: { @@ -8722,11 +8776,11 @@ 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) { +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___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) { _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); +x_11 = l_MonadExcept_ofExcept___at_Lean_Elab_Tactic_evalTactic___spec__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -8865,7 +8919,7 @@ x_15 = l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3(x_1, x_2, 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) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___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, 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; @@ -8873,16 +8927,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__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); +x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_evalTactic___spec__8(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__10___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___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) { _start: { lean_object* x_11; -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); +x_11 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_evalTactic___spec__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -8894,11 +8948,11 @@ lean_dec(x_2); return x_11; } } -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) { +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Tactic_evalTactic___spec__10___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__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); +x_14 = l_Lean_profileitM___at_Lean_Elab_Tactic_evalTactic___spec__10___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; diff --git a/stage0/stdlib/Lean/Elab/Term.c b/stage0/stdlib/Lean/Elab/Term.c index 5655d8f47189..1e80009f086b 100644 --- a/stage0/stdlib/Lean/Elab/Term.c +++ b/stage0/stdlib/Lean/Elab/Term.c @@ -13,7 +13,6 @@ #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_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; LEAN_EXPORT lean_object* l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSavedPartialInfoContext___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__4___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -31,6 +30,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_withoutSavingRecAppSyntax___rarg(lean_ static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_postponeElabTermCore___closed__4; 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; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__6; 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_withoutHeedElabAsElimImp(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*); @@ -84,9 +84,9 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_withDeclName static lean_object* l_Lean_Elab_Term_withExpectedType___closed__2; static lean_object* l_Lean_Elab_Term_TermElabM_toIO___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_withoutAutoBoundImplicit(lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__4; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Term_0__Lean_Elab_Term_withoutModifyingStateWithInfoAndMessagesImpl___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*); static lean_object* l_Lean_Elab_Term_ensureType___lambda__1___closed__3; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__4; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__35(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_replace___at_Lean_Elab_Term_ContainsPendingMVar_visit___spec__6(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkFreshLevelMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -134,7 +134,6 @@ static lean_object* l_Lean_Elab_Term_instMetaEvalTermElabM___rarg___closed__24; LEAN_EXPORT lean_object* l_Lean_Elab_Term_registerSyntheticMVar___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_registerSyntheticMVarWithCurrRef(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__4(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_initFn____x40_Lean_Elab_Term___hyg_16089____closed__7; LEAN_EXPORT lean_object* l_Lean_Elab_Term_withLevelNames(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__34___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_tryPostponeIfHasMVars___closed__4; @@ -229,10 +228,10 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_mkConst___spec__2 static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_isNoImplicitLambda___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveLocalName(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__1; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__10; lean_object* l_Lean_mkHashSetImp___rarg(lean_object*); lean_object* l_Lean_Core_getMessageLog___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_expandDeclId___closed__1; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__14; lean_object* l_Lean_Core_checkSystem(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Elab_Term_resolveId_x3f___spec__3(lean_object*, lean_object*); lean_object* l_Lean_instBEqLocalInstance___boxed(lean_object*, lean_object*); @@ -258,6 +257,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_TermElabM_toIO___rarg(lean_object*, le static lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Term_expandDeclId___spec__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_CommandContextInfo_saveNoFileMap___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__3(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_mkTermElabAttributeUnsafe___closed__4; +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Term_mkCoe___spec__5(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___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_instMonadTermElabM; LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__5(lean_object*, lean_object*); @@ -341,6 +341,7 @@ LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Elab_Term_getMVarErrorInfo 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*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_saveContext___boxed(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_13906____closed__11; lean_object* l_Lean_MessageData_hasSyntheticSorry(lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMapImp_expand___at_Lean_Elab_Term_ContainsPendingMVar_visit___spec__3(lean_object*, lean_object*); lean_object* l_Lean_mkIdentFrom(lean_object*, lean_object*, uint8_t); @@ -348,7 +349,6 @@ LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_Elab_Term_0__Lean_Ela static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_dropTermParens___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_withoutHeedElabAsElim___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_instMonadBacktrackSavedStateTermElabM___closed__3; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__5; static lean_object* l_Lean_Elab_Term_instMetaEvalTermElabM___rarg___closed__21; 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*); @@ -367,6 +367,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_throwAbortTerm___at_Lean_Elab_Term_throwMVa 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; +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____closed__8; lean_object* lean_io_get_num_heartbeats(lean_object*); uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); lean_object* l_Lean_Syntax_identComponents(lean_object*, lean_object*); @@ -387,10 +388,11 @@ lean_object* l_Lean_Expr_mvar___override(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_State_syntheticMVars___default; static lean_object* l_Lean_Elab_Term_instMetaEvalTermElabM___rarg___closed__25; +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____closed__6; static lean_object* l_Lean_Elab_Term_MVarErrorInfo_logError___closed__9; uint8_t l___private_Lean_Util_Trace_0__Lean_checkTraceOption(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__13; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Cache_post___default; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__2; static lean_object* l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__3; static lean_object* l_Lean_Elab_Term_MVarErrorInfo_logError___closed__10; LEAN_EXPORT lean_object* l_Lean_Elab_Term_withoutHeedElabAsElim___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -416,6 +418,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_isLambdaWi 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); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__7; LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveLocalName_loop___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*); static lean_object* l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__10; static lean_object* l_Lean_Elab_throwAutoBoundImplicitLocal___at_Lean_Elab_Term_resolveName_process___spec__1___closed__1; @@ -438,16 +441,15 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_isLetRecAuxMVar___lambda__3(lean_objec lean_object* l_instBEqProd___rarg(lean_object*, lean_object*, lean_object*, 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*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16089____closed__1; 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*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__13; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallBoundedTelescope___at_Lean_Elab_Term_addAutoBoundImplicits_x27___spec__2(lean_object*); static lean_object* l_Lean_Elab_Term_MVarErrorInfo_logError___closed__12; LEAN_EXPORT lean_object* l_Array_findSomeRevM_x3f_find___at_Lean_Elab_Term_resolveLocalName___spec__11(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Elab_Term_resolveName___spec__2___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____closed__5; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__19___boxed(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__28(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo(lean_object*); @@ -497,11 +499,13 @@ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_addAutoBound lean_object* l_IO_println___at_Lean_instEval___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Elab_abortCommandExceptionId; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__8; lean_object* l_List_appendTR___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_instInhabitedSyntheticMVarDecl___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___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_Elab_Term_withoutHeedElabAsElimImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_instApplicativeReaderT___rarg(lean_object*); +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Term_mkCoe___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_instMonadMacroAdapterTermElabM___closed__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; @@ -518,7 +522,6 @@ static lean_object* l_Lean_Elab_Term_mkTypeMismatchError___closed__2; lean_object* l_Lean_Meta_mkFreshLevelMVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___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_withoutPostponing___rarg(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_16089____closed__4; static lean_object* l_Lean_Elab_Term_getFVarLocalDecl_x21___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at___private_Lean_Elab_Term_0__Lean_Elab_Term_withoutModifyingStateWithInfoAndMessagesImpl___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_instMonadTermElabM___closed__4; @@ -529,6 +532,7 @@ static lean_object* l_Lean_Elab_Term_isLetRecAuxMVar___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveName___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_throwPostpone___at_Lean_Elab_Term_tryPostpone___spec__1___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____closed__7; 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*); @@ -537,6 +541,7 @@ 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; static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_isExplicitApp___closed__2; static lean_object* l_Lean_Elab_Term_instToStringSyntheticMVarKind___closed__4; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__14; 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_withoutHeedElabAsElim___rarg___closed__1; @@ -625,6 +630,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabUsingE 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_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__2; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__51___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_TermElabM_toIO(lean_object*); lean_object* l_Lean_Meta_SavedState_restore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -652,6 +658,7 @@ static lean_object* l_Lean_Elab_Term_instMetaEvalTermElabM___rarg___closed__9; LEAN_EXPORT lean_object* l_Lean_Elab_Term_registerSyntheticMVarWithCurrRef___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_PersistentArray_anyM___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__29___boxed(lean_object*, lean_object*); 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_initFn____x40_Lean_Elab_Term___hyg_13906_(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; lean_object* l_Lean_Syntax_getKind(lean_object*); @@ -660,6 +667,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_isTacticOrPostponedHole_x3f(lean_objec uint8_t l_Lean_NameMap_contains___rarg(lean_object*, lean_object*); lean_object* l_Lean_PersistentArray_append___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_LVal_getRef(lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__9; lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); extern lean_object* l___private_Lean_DocString_0__Lean_docStringExt; uint8_t l_Lean_Expr_isMVar(lean_object*); @@ -680,7 +688,6 @@ 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*); 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*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__10; 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; @@ -690,6 +697,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_mkSyntheti lean_object* l_Lean_Name_replacePrefix(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkTermInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Context_autoBoundImplicits___default___closed__2; +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____closed__9; static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabUsingElabFnsAux___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_useImplicitLambda___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_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___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*); @@ -711,7 +719,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkFreshBinderName___rarg(lean_object*, 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*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16089____closed__2; 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___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*); @@ -756,7 +763,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_getFVarLocalDecl_x21(lean_object*, lea LEAN_EXPORT lean_object* l_Lean_Elab_Term_getDeclName_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_coerceToSort_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_liftLevelM___rarg___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_16089____closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_withSaveInfoContext___at___private_Lean_Elab_Term_0__Lean_Elab_Term_withoutModifyingStateWithInfoAndMessagesImpl___spec__1(lean_object*); static lean_object* l_Lean_Elab_Term_instMetaEvalTermElabM___rarg___closed__5; uint8_t l___private_Lean_Attributes_0__Lean_beqAttributeApplicationTime____x40_Lean_Attributes___hyg_15_(uint8_t, uint8_t); @@ -930,7 +936,6 @@ 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_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*); @@ -964,9 +969,7 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_instMetaE LEAN_EXPORT 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_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_initFn____x40_Lean_Elab_Term___hyg_13831____closed__8; static lean_object* l_Lean_Elab_Term_instMetaEvalTermElabM___rarg___closed__4; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__7; 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*); @@ -982,16 +985,16 @@ LEAN_EXPORT lean_object* l_Lean_profileitM___at___private_Lean_Elab_Term_0__Lean 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*); -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16089_(lean_object*); static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_isTypeAscription___closed__1; +lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNode___spec__1(size_t, size_t, 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*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkFreshIdent___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_getDeclName_x3f___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_16164____closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_withDeclName___spec__12(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_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; @@ -1042,7 +1045,6 @@ 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*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__22(lean_object*, lean_object*); static lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Term_expandDeclId___spec__6___lambda__1___closed__4; -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16089____closed__3; static lean_object* l_Lean_Elab_Term_MVarErrorInfo_logError___closed__13; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Context_tacticCache_x3f___default; LEAN_EXPORT lean_object* l_Lean_Elab_Term_universeConstraintsCheckpoint___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1056,10 +1058,11 @@ static lean_object* l_Lean_Elab_Term_MVarErrorInfo_logError_addArgName___closed_ LEAN_EXPORT lean_object* l_Lean_Elab_Term_saveState(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_tryPostponeIfMVar___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Elab_Term_getSyntheticMVarDecl_x3f___spec__1___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__11; 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_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__15; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__12; 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*); 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; @@ -1109,6 +1112,7 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Term_ LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__7(lean_object*, lean_object*, size_t, size_t); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withAutoBoundImplicit(lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____closed__3; static lean_object* l_Lean_Elab_logException___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__7___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_getMVarDecl___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*); @@ -1170,7 +1174,6 @@ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_logUnassigne 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*); static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_1977____closed__2; 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_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__15; 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*); @@ -1208,7 +1211,6 @@ static lean_object* l_Lean_Elab_Term_instMetaEvalTermElabM___rarg___closed__15; LEAN_EXPORT lean_object* l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSavedPartialInfoContext___at___private_Lean_Elab_Term_0__Lean_Elab_Term_withoutModifyingStateWithInfoAndMessagesImpl___spec__2___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_contains___at_Lean_Elab_Term_logUnassignedUsingErrorInfos___spec__1___boxed(lean_object*, lean_object*); 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*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__3; 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; @@ -1312,6 +1314,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*); LEAN_EXPORT lean_object* l_Lean_Elab_withSaveParentDeclInfoContext___at_Lean_Elab_Term_withDeclName___spec__1(lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____closed__2; 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; @@ -1343,7 +1346,6 @@ static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___l 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*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSavedPartialInfoContext___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__4___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_Elab_Term_withLevelNames___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_mkExplicitBinder___closed__2; @@ -1369,7 +1371,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_addAutoBoundImplicits_go___lambda__1__ lean_object* lean_io_error_to_string(lean_object*); lean_object* l_Lean_KeyedDeclsAttribute_getEntries___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT 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*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__12; static lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1___closed__3; extern lean_object* l_Lean_instInhabitedSyntax; LEAN_EXPORT lean_object* l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1380,8 +1381,8 @@ LEAN_EXPORT lean_object* l_Lean_Elab_CommandContextInfo_saveNoFileMap___at_Lean_ 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*); lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__1; 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*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____closed__4; 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*); 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_EXPORT lean_object* l_Lean_localDeclDependsOn___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1402,6 +1403,7 @@ lean_object* lean_string_append(lean_object*, lean_object*); 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*); 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*); static lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Term_expandDeclId___spec__6___closed__1; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__3; static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__1___closed__3; LEAN_EXPORT uint8_t l_Array_contains___at_Lean_Elab_Term_logUnassignedUsingErrorInfos___spec__1(lean_object*, lean_object*); lean_object* l_Lean_getAttributeImpl(lean_object*, lean_object*); @@ -1409,6 +1411,7 @@ LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Elab_Term_instAddErrorMessageCon LEAN_EXPORT lean_object* l_Lean_Elab_Term_withoutErrToSorry(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_State_pendingMVars___default; static lean_object* l_List_foldl___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__2___closed__3; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__5; lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_Term_instInhabitedLetRecToLift___closed__4; static lean_object* l_Lean_Elab_Term_mkExplicitBinder___closed__7; @@ -1416,7 +1419,6 @@ 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; @@ -1437,10 +1439,10 @@ static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_isLambdaWithImp extern lean_object* l_Lean_Elab_unsupportedSyntaxExceptionId; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__16(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__9; uint8_t lean_usize_dec_lt(size_t, size_t); LEAN_EXPORT uint8_t l_Lean_Elab_Term_Context_errToSorry___default; LEAN_EXPORT lean_object* l_Lean_Elab_Term_withAuxDecl___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_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__1; LEAN_EXPORT lean_object* l_Array_findSomeRevM_x3f_find___at_Lean_Elab_Term_resolveLocalName___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeRevMAux___at_Lean_Elab_Term_resolveLocalName___spec__9(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveId_x3f___lambda__2(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1457,7 +1459,6 @@ static lean_object* l_Lean_Elab_Term_instMetaEvalTermElabM___rarg___closed__19; static lean_object* l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__9; LEAN_EXPORT lean_object* l_Lean_Elab_Term_withAutoBoundImplicitForbiddenPred___rarg___lambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_instMetaEvalTermElabM___rarg___closed__13; -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16089____closed__9; lean_object* lean_nat_add(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_instToStringSyntheticMVarKind___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_withMacroExpansionInfo___at_Lean_Elab_Term_withMacroExpansion___spec__1(lean_object*); @@ -1529,6 +1530,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Term_0_ static lean_object* l_Lean_Elab_Term_exprToSyntax___lambda__1___closed__3; static lean_object* l_Lean_Elab_Term_instMetaEvalTermElabM___rarg___closed__26; LEAN_EXPORT 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_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164_(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__10___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_throwError___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Term_0__Lean_Elab_Term_withoutModifyingStateWithInfoAndMessagesImpl___spec__7(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1551,16 +1553,12 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveName_x27(lean_object*, lean_obj LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at_Lean_Elab_Term_ContainsPendingMVar_visit___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabType(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_16089____closed__6; LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Elab_Term_resolveName_x27___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___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_16089____closed__8; 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*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__6; 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*); @@ -26921,125 +26919,6 @@ return x_40; } } } -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_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_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_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_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; -} -} -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_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_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: { @@ -27047,76 +26926,269 @@ 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; +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; size_t x_22; size_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; 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_16 = lean_st_ref_get(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); +x_19 = lean_ctor_get(x_17, 3); +lean_inc(x_19); +lean_dec(x_17); +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_addTraceNode___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 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_25, x_8, x_9, x_10, x_11, x_18); lean_dec(x_10); -return x_19; +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_st_ref_take(x_11, 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_is_exclusive(x_30); +if (x_32 == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_33 = lean_ctor_get(x_30, 3); +lean_dec(x_33); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_3); +lean_ctor_set(x_34, 1, x_27); +x_35 = l_Lean_PersistentArray_push___rarg(x_1, x_34); +lean_ctor_set(x_30, 3, x_35); +x_36 = lean_st_ref_set(x_11, x_30, x_31); +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 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; } 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; uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -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); -x_31 = lean_ctor_get_uint8(x_10, sizeof(void*)*11); -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_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; +} +} +else +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_43 = lean_ctor_get(x_30, 0); +x_44 = lean_ctor_get(x_30, 1); +x_45 = lean_ctor_get(x_30, 2); +x_46 = lean_ctor_get(x_30, 4); +x_47 = lean_ctor_get(x_30, 5); +x_48 = lean_ctor_get(x_30, 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_dec(x_30); +x_49 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_49, 0, x_3); +lean_ctor_set(x_49, 1, x_27); +x_50 = l_Lean_PersistentArray_push___rarg(x_1, x_49); +x_51 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_51, 0, x_43); +lean_ctor_set(x_51, 1, x_44); +lean_ctor_set(x_51, 2, x_45); +lean_ctor_set(x_51, 3, x_50); +lean_ctor_set(x_51, 4, x_46); +lean_ctor_set(x_51, 5, x_47); +lean_ctor_set(x_51, 6, x_48); +x_52 = lean_st_ref_set(x_11, x_51, x_31); +x_53 = lean_ctor_get(x_52, 1); +lean_inc(x_53); +if (lean_is_exclusive(x_52)) { + lean_ctor_release(x_52, 0); + lean_ctor_release(x_52, 1); + x_54 = x_52; +} else { + lean_dec_ref(x_52); + x_54 = lean_box(0); +} +x_55 = lean_box(0); +if (lean_is_scalar(x_54)) { + x_56 = lean_alloc_ctor(0, 2, 0); +} else { + x_56 = x_54; +} +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_53); +return x_56; +} +} +else +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; uint8_t x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; size_t x_77; size_t x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; +x_57 = lean_ctor_get(x_10, 0); +x_58 = lean_ctor_get(x_10, 1); +x_59 = lean_ctor_get(x_10, 2); +x_60 = lean_ctor_get(x_10, 3); +x_61 = lean_ctor_get(x_10, 4); +x_62 = lean_ctor_get(x_10, 5); +x_63 = lean_ctor_get(x_10, 6); +x_64 = lean_ctor_get(x_10, 7); +x_65 = lean_ctor_get(x_10, 8); +x_66 = lean_ctor_get(x_10, 9); +x_67 = lean_ctor_get(x_10, 10); +x_68 = lean_ctor_get_uint8(x_10, sizeof(void*)*11); +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_dec(x_10); -x_32 = l_Lean_replaceRef(x_3, x_25); -lean_dec(x_25); -x_33 = lean_alloc_ctor(0, 11, 1); -lean_ctor_set(x_33, 0, x_20); -lean_ctor_set(x_33, 1, x_21); -lean_ctor_set(x_33, 2, x_22); -lean_ctor_set(x_33, 3, x_23); -lean_ctor_set(x_33, 4, x_24); -lean_ctor_set(x_33, 5, x_32); -lean_ctor_set(x_33, 6, x_26); -lean_ctor_set(x_33, 7, x_27); -lean_ctor_set(x_33, 8, x_28); -lean_ctor_set(x_33, 9, x_29); -lean_ctor_set(x_33, 10, x_30); -lean_ctor_set_uint8(x_33, sizeof(void*)*11, x_31); -x_34 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_4, x_8, x_9, x_33, x_11, x_12); -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___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Elab_Term_mkCoe___spec__5(x_1, x_2, x_3, x_35, x_5, x_6, x_7, x_8, x_9, x_33, x_11, x_36); -lean_dec(x_33); -return x_37; +x_69 = l_Lean_replaceRef(x_3, x_62); +lean_dec(x_62); +x_70 = lean_alloc_ctor(0, 11, 1); +lean_ctor_set(x_70, 0, x_57); +lean_ctor_set(x_70, 1, x_58); +lean_ctor_set(x_70, 2, x_59); +lean_ctor_set(x_70, 3, x_60); +lean_ctor_set(x_70, 4, x_61); +lean_ctor_set(x_70, 5, x_69); +lean_ctor_set(x_70, 6, x_63); +lean_ctor_set(x_70, 7, x_64); +lean_ctor_set(x_70, 8, x_65); +lean_ctor_set(x_70, 9, x_66); +lean_ctor_set(x_70, 10, x_67); +lean_ctor_set_uint8(x_70, sizeof(void*)*11, x_68); +x_71 = lean_st_ref_get(x_11, x_12); +x_72 = lean_ctor_get(x_71, 0); +lean_inc(x_72); +x_73 = lean_ctor_get(x_71, 1); +lean_inc(x_73); +lean_dec(x_71); +x_74 = lean_ctor_get(x_72, 3); +lean_inc(x_74); +lean_dec(x_72); +x_75 = l_Lean_PersistentArray_toArray___rarg(x_74); +x_76 = lean_array_get_size(x_75); +x_77 = lean_usize_of_nat(x_76); +lean_dec(x_76); +x_78 = 0; +x_79 = l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNode___spec__1(x_77, x_78, x_75); +x_80 = lean_alloc_ctor(9, 3, 1); +lean_ctor_set(x_80, 0, x_2); +lean_ctor_set(x_80, 1, x_4); +lean_ctor_set(x_80, 2, x_79); +lean_ctor_set_uint8(x_80, sizeof(void*)*3, x_5); +x_81 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_80, x_8, x_9, x_70, x_11, x_73); +lean_dec(x_70); +x_82 = lean_ctor_get(x_81, 0); +lean_inc(x_82); +x_83 = lean_ctor_get(x_81, 1); +lean_inc(x_83); +lean_dec(x_81); +x_84 = lean_st_ref_take(x_11, x_83); +x_85 = lean_ctor_get(x_84, 0); +lean_inc(x_85); +x_86 = lean_ctor_get(x_84, 1); +lean_inc(x_86); +lean_dec(x_84); +x_87 = lean_ctor_get(x_85, 0); +lean_inc(x_87); +x_88 = lean_ctor_get(x_85, 1); +lean_inc(x_88); +x_89 = lean_ctor_get(x_85, 2); +lean_inc(x_89); +x_90 = lean_ctor_get(x_85, 4); +lean_inc(x_90); +x_91 = lean_ctor_get(x_85, 5); +lean_inc(x_91); +x_92 = lean_ctor_get(x_85, 6); +lean_inc(x_92); +if (lean_is_exclusive(x_85)) { + lean_ctor_release(x_85, 0); + lean_ctor_release(x_85, 1); + lean_ctor_release(x_85, 2); + lean_ctor_release(x_85, 3); + lean_ctor_release(x_85, 4); + lean_ctor_release(x_85, 5); + lean_ctor_release(x_85, 6); + x_93 = x_85; +} else { + lean_dec_ref(x_85); + x_93 = lean_box(0); } +x_94 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_94, 0, x_3); +lean_ctor_set(x_94, 1, x_82); +x_95 = l_Lean_PersistentArray_push___rarg(x_1, x_94); +if (lean_is_scalar(x_93)) { + x_96 = lean_alloc_ctor(0, 7, 0); +} else { + x_96 = x_93; } +lean_ctor_set(x_96, 0, x_87); +lean_ctor_set(x_96, 1, x_88); +lean_ctor_set(x_96, 2, x_89); +lean_ctor_set(x_96, 3, x_95); +lean_ctor_set(x_96, 4, x_90); +lean_ctor_set(x_96, 5, x_91); +lean_ctor_set(x_96, 6, x_92); +x_97 = lean_st_ref_set(x_11, x_96, x_86); +x_98 = lean_ctor_get(x_97, 1); +lean_inc(x_98); +if (lean_is_exclusive(x_97)) { + lean_ctor_release(x_97, 0); + lean_ctor_release(x_97, 1); + x_99 = x_97; +} else { + lean_dec_ref(x_97); + x_99 = lean_box(0); } -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) { +x_100 = lean_box(0); +if (lean_is_scalar(x_99)) { + x_101 = lean_alloc_ctor(0, 2, 0); +} else { + x_101 = x_99; +} +lean_ctor_set(x_101, 0, x_100); +lean_ctor_set(x_101, 1, x_98); +return x_101; +} +} +} +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Term_mkCoe___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: { if (lean_obj_tag(x_1) == 0) @@ -27216,7 +27288,7 @@ x_15 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Term_mkCo 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); +x_17 = l_MonadExcept_ofExcept___at_Lean_Elab_Term_mkCoe___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); return x_17; @@ -27491,7 +27563,7 @@ x_32 = lean_st_ref_set(x_14, x_27, x_28); x_33 = lean_ctor_get(x_32, 1); lean_inc(x_33); lean_dec(x_32); -x_34 = l_MonadExcept_ofExcept___at_Lean_Elab_Term_mkCoe___spec__6(x_23, x_9, x_10, x_11, x_12, x_13, x_14, x_33); +x_34 = l_MonadExcept_ofExcept___at_Lean_Elab_Term_mkCoe___spec__5(x_23, x_9, x_10, x_11, x_12, x_13, x_14, x_33); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); @@ -27575,7 +27647,7 @@ x_52 = lean_st_ref_set(x_14, x_51, x_28); x_53 = lean_ctor_get(x_52, 1); lean_inc(x_53); lean_dec(x_52); -x_54 = l_MonadExcept_ofExcept___at_Lean_Elab_Term_mkCoe___spec__6(x_23, x_9, x_10, x_11, x_12, x_13, x_14, x_53); +x_54 = l_MonadExcept_ofExcept___at_Lean_Elab_Term_mkCoe___spec__5(x_23, x_9, x_10, x_11, x_12, x_13, x_14, x_53); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); @@ -28747,22 +28819,6 @@ lean_dec(x_1); return x_6; } } -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: -{ -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_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); -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_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: { @@ -28778,11 +28834,11 @@ lean_dec(x_6); return x_14; } } -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) { +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___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) { _start: { 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); +x_9 = l_MonadExcept_ofExcept___at_Lean_Elab_Term_mkCoe___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); @@ -43316,75 +43372,290 @@ lean_inc(x_1); x_13 = l_Lean_Elab_Term_elabTerm(x_1, x_2, x_3, x_4, x_6, x_7, x_8, x_9, 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_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; uint8_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; x_14 = lean_ctor_get(x_13, 0); lean_inc(x_14); x_15 = lean_ctor_get(x_13, 1); lean_inc(x_15); lean_dec(x_13); x_16 = lean_box(0); -x_17 = !lean_is_exclusive(x_10); -if (x_17 == 0) -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_10, 5); -x_19 = l_Lean_replaceRef(x_1, x_18); -lean_dec(x_18); +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, 3); +lean_inc(x_20); +x_21 = lean_ctor_get(x_10, 4); +lean_inc(x_21); +x_22 = lean_ctor_get(x_10, 5); +lean_inc(x_22); +x_23 = lean_ctor_get(x_10, 6); +lean_inc(x_23); +x_24 = lean_ctor_get(x_10, 7); +lean_inc(x_24); +x_25 = lean_ctor_get(x_10, 8); +lean_inc(x_25); +x_26 = lean_ctor_get(x_10, 9); +lean_inc(x_26); +x_27 = lean_ctor_get(x_10, 10); +lean_inc(x_27); +x_28 = lean_ctor_get_uint8(x_10, sizeof(void*)*11); +x_29 = l_Lean_replaceRef(x_1, x_22); +lean_dec(x_22); lean_dec(x_1); -lean_ctor_set(x_10, 5, x_19); -x_20 = l_Lean_Elab_Term_ensureHasType(x_2, x_14, x_5, x_16, x_6, x_7, x_8, x_9, x_10, x_11, x_15); -return x_20; +x_30 = lean_alloc_ctor(0, 11, 1); +lean_ctor_set(x_30, 0, x_17); +lean_ctor_set(x_30, 1, x_18); +lean_ctor_set(x_30, 2, x_19); +lean_ctor_set(x_30, 3, x_20); +lean_ctor_set(x_30, 4, x_21); +lean_ctor_set(x_30, 5, x_29); +lean_ctor_set(x_30, 6, x_23); +lean_ctor_set(x_30, 7, x_24); +lean_ctor_set(x_30, 8, x_25); +lean_ctor_set(x_30, 9, x_26); +lean_ctor_set(x_30, 10, x_27); +lean_ctor_set_uint8(x_30, sizeof(void*)*11, x_28); +lean_inc(x_11); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +x_31 = l_Lean_Elab_Term_ensureHasType(x_2, x_14, x_5, x_16, x_6, x_7, x_8, x_9, x_30, x_11, x_15); +if (lean_obj_tag(x_31) == 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_2); +return x_31; } 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; uint8_t x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_21 = lean_ctor_get(x_10, 0); -x_22 = lean_ctor_get(x_10, 1); -x_23 = lean_ctor_get(x_10, 2); -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, 6); -x_28 = lean_ctor_get(x_10, 7); -x_29 = lean_ctor_get(x_10, 8); -x_30 = lean_ctor_get(x_10, 9); -x_31 = lean_ctor_get(x_10, 10); -x_32 = lean_ctor_get_uint8(x_10, sizeof(void*)*11); -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_inc(x_21); +uint8_t x_32; +x_32 = !lean_is_exclusive(x_31); +if (x_32 == 0) +{ +lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_33 = lean_ctor_get(x_31, 0); +x_34 = lean_ctor_get(x_31, 1); +x_35 = l_Lean_Exception_isRuntime(x_33); +if (x_35 == 0) +{ +uint8_t x_36; +x_36 = lean_ctor_get_uint8(x_6, sizeof(void*)*8 + 1); +if (x_36 == 0) +{ +lean_dec(x_11); lean_dec(x_10); -x_33 = l_Lean_replaceRef(x_1, x_26); -lean_dec(x_26); -lean_dec(x_1); -x_34 = lean_alloc_ctor(0, 11, 1); -lean_ctor_set(x_34, 0, x_21); -lean_ctor_set(x_34, 1, x_22); -lean_ctor_set(x_34, 2, x_23); -lean_ctor_set(x_34, 3, x_24); -lean_ctor_set(x_34, 4, x_25); -lean_ctor_set(x_34, 5, x_33); -lean_ctor_set(x_34, 6, x_27); -lean_ctor_set(x_34, 7, x_28); -lean_ctor_set(x_34, 8, x_29); -lean_ctor_set(x_34, 9, x_30); -lean_ctor_set(x_34, 10, x_31); -lean_ctor_set_uint8(x_34, sizeof(void*)*11, x_32); -x_35 = l_Lean_Elab_Term_ensureHasType(x_2, x_14, x_5, x_16, x_6, x_7, x_8, x_9, x_34, x_11, x_15); -return x_35; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +return x_31; } +else +{ +if (lean_obj_tag(x_33) == 0) +{ +lean_object* x_37; +lean_free_object(x_31); +x_37 = l_Lean_Elab_Term_exceptionToSorry(x_33, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_34); +return x_37; +} +else +{ +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); +return x_31; +} +} +} +else +{ +if (x_28 == 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_2); +return x_31; +} +else +{ +uint8_t x_38; +x_38 = lean_ctor_get_uint8(x_6, sizeof(void*)*8 + 1); +if (x_38 == 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_2); +return x_31; +} +else +{ +if (lean_obj_tag(x_33) == 0) +{ +lean_object* x_39; +lean_free_object(x_31); +x_39 = l_Lean_Elab_Term_exceptionToSorry(x_33, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_34); +return x_39; } else { -uint8_t x_36; +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); +return x_31; +} +} +} +} +} +else +{ +lean_object* x_40; lean_object* x_41; uint8_t x_42; +x_40 = lean_ctor_get(x_31, 0); +x_41 = lean_ctor_get(x_31, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_31); +x_42 = l_Lean_Exception_isRuntime(x_40); +if (x_42 == 0) +{ +uint8_t x_43; +x_43 = lean_ctor_get_uint8(x_6, sizeof(void*)*8 + 1); +if (x_43 == 0) +{ +lean_object* x_44; +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_44 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_44, 0, x_40); +lean_ctor_set(x_44, 1, x_41); +return x_44; +} +else +{ +if (lean_obj_tag(x_40) == 0) +{ +lean_object* x_45; +x_45 = l_Lean_Elab_Term_exceptionToSorry(x_40, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_41); +return x_45; +} +else +{ +lean_object* x_46; +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_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_40); +lean_ctor_set(x_46, 1, x_41); +return x_46; +} +} +} +else +{ +if (x_28 == 0) +{ +lean_object* x_47; +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_47 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_47, 0, x_40); +lean_ctor_set(x_47, 1, x_41); +return x_47; +} +else +{ +uint8_t x_48; +x_48 = lean_ctor_get_uint8(x_6, sizeof(void*)*8 + 1); +if (x_48 == 0) +{ +lean_object* x_49; +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_49 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_49, 0, x_40); +lean_ctor_set(x_49, 1, x_41); +return x_49; +} +else +{ +if (lean_obj_tag(x_40) == 0) +{ +lean_object* x_50; +x_50 = l_Lean_Elab_Term_exceptionToSorry(x_40, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_41); +return x_50; +} +else +{ +lean_object* x_51; +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_51 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_51, 0, x_40); +lean_ctor_set(x_51, 1, x_41); +return x_51; +} +} +} +} +} +} +} +else +{ +uint8_t x_52; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -43394,23 +43665,23 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); -x_36 = !lean_is_exclusive(x_13); -if (x_36 == 0) +x_52 = !lean_is_exclusive(x_13); +if (x_52 == 0) { return x_13; } 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_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_13, 0); +x_54 = lean_ctor_get(x_13, 1); +lean_inc(x_54); +lean_inc(x_53); 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; +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; } } } @@ -51943,7 +52214,7 @@ lean_dec(x_3); return x_10; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__1() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__1() { _start: { lean_object* x_1; @@ -51951,17 +52222,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_13831____closed__2() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____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_13831____closed__1; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____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_13831____closed__3() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -51971,27 +52242,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_13831____closed__4() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____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_13831____closed__3; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____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_13831____closed__5() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____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_13831____closed__4; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____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_13831____closed__6() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__6() { _start: { lean_object* x_1; @@ -51999,17 +52270,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_13831____closed__7() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____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_13831____closed__5; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__6; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__5; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____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_13831____closed__8() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__8() { _start: { lean_object* x_1; @@ -52017,47 +52288,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_13831____closed__9() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____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_13831____closed__7; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__8; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__7; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____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_13831____closed__10() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____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_13831____closed__9; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____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_13831____closed__11() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____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_13831____closed__10; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____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_13831____closed__12() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____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_13831____closed__11; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____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_13831____closed__13() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__13() { _start: { lean_object* x_1; @@ -52065,33 +52336,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_13831____closed__14() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____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_13831____closed__12; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__13; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__12; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____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_13831____closed__15() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____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_13831____closed__14; -x_2 = lean_unsigned_to_nat(13831u); +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__14; +x_2 = lean_unsigned_to_nat(13906u); 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_13831_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906_(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_13831____closed__2; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__2; x_3 = 0; -x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__15; +x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__15; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -52418,7 +52689,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_13831____closed__2; +x_9 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____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); @@ -60799,77 +61070,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_16089____closed__1() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____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_13831____closed__4; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__6; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__4; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____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_16089____closed__2() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____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_16089____closed__1; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__8; +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____closed__1; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____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_16089____closed__3() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____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_16089____closed__2; +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____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_16089____closed__4() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____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_16089____closed__3; +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____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_16089____closed__5() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____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_16089____closed__4; +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____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_16089____closed__6() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____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_16089____closed__5; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__13; +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____closed__5; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____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_16089____closed__7() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____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_16089____closed__6; -x_2 = lean_unsigned_to_nat(16089u); +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____closed__6; +x_2 = lean_unsigned_to_nat(16164u); 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_16089____closed__8() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____closed__8() { _start: { lean_object* x_1; @@ -60877,23 +61148,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_16089____closed__9() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____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_16089____closed__8; +x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____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_16089_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164_(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_16089____closed__7; +x_4 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____closed__7; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); if (lean_obj_tag(x_5) == 0) { @@ -60909,7 +61180,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_16089____closed__9; +x_10 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____closed__9; x_11 = l_Lean_registerTraceClass(x_10, x_3, x_4, x_9); return x_11; } @@ -61589,37 +61860,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_13831____closed__1 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__1(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__1); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__2 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__2(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__2); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__3 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__3(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__3); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__4 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__4(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__4); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__5 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__5(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__5); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__6 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__6(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__6); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__7 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__7(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__7); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__8 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__8(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__8); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__9 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__9(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__9); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__10 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__10(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__10); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__11 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__11(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__11); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__12 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__12(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__12); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__13 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__13(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__13); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__14 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__14(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__14); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__15 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__15(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831____closed__15); -if (builtin) {res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13831_(lean_io_mk_world()); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__1 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__1); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__2 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__2(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__2); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__3 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__3(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__3); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__4 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__4(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__4); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__5 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__5(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__5); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__6 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__6(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__6); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__7 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__7(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__7); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__8 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__8(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__8); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__9 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__9(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__9); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__10 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__10(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__10); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__11 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__11(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__11); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__12 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__12(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__12); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__13 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__13(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__13); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__14 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__14(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__14); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__15 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__15(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906____closed__15); +if (builtin) {res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13906_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }l_Lean_Elab_Term_isLetRecAuxMVar___lambda__3___closed__1 = _init_l_Lean_Elab_Term_isLetRecAuxMVar___lambda__3___closed__1(); @@ -61836,25 +62107,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_16089____closed__1 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16089____closed__1(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16089____closed__1); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16089____closed__2 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16089____closed__2(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16089____closed__2); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16089____closed__3 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16089____closed__3(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16089____closed__3); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16089____closed__4 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16089____closed__4(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16089____closed__4); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16089____closed__5 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16089____closed__5(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16089____closed__5); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16089____closed__6 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16089____closed__6(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16089____closed__6); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16089____closed__7 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16089____closed__7(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16089____closed__7); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16089____closed__8 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16089____closed__8(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16089____closed__8); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16089____closed__9 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16089____closed__9(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16089____closed__9); -if (builtin) {res = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16089_(lean_io_mk_world()); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____closed__1 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____closed__1(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____closed__1); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____closed__2 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____closed__2(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____closed__2); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____closed__3 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____closed__3(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____closed__3); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____closed__4 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____closed__4(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____closed__4); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____closed__5 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____closed__5(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____closed__5); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____closed__6 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____closed__6(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____closed__6); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____closed__7 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____closed__7(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____closed__7); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____closed__8 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____closed__8(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____closed__8); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____closed__9 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____closed__9(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164____closed__9); +if (builtin) {res = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_16164_(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/Basic.c b/stage0/stdlib/Lean/Meta/Basic.c index 9c97950df61b..dfd885185ae8 100644 --- a/stage0/stdlib/Lean/Meta/Basic.c +++ b/stage0/stdlib/Lean/Meta/Basic.c @@ -15,7 +15,6 @@ extern "C" { #endif LEAN_EXPORT lean_object* l_Lean_Meta_isDefEqStuckExceptionId; LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withNewLocalInstancesImpAux___rarg(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*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_exposeRelevantUniverses(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_instMonadMetaM___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_processPostponed_loop___closed__1; @@ -455,7 +454,6 @@ 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*); 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*); @@ -658,7 +656,6 @@ 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); 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*); @@ -669,6 +666,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getP LEAN_EXPORT lean_object* l_Lean_Meta_getConfig(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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_object* l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNode___spec__1(size_t, size_t, 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*); LEAN_EXPORT lean_object* l_Lean_Meta_mapForallTelescope___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -892,7 +890,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_instInhabitedMetaM___rarg(lean_object*); static lean_object* l_Lean_Meta_instMonadMetaM___closed__6; 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); @@ -955,7 +952,6 @@ 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*); @@ -974,6 +970,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_isClassImp_x3f static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1614____closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_useEtaStruct___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_findMVarDecl_x3f___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__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_withLocalDeclsD___spec__1___rarg(size_t, size_t, lean_object*); static lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__5___closed__1; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_liftMkBindingM___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1051,6 +1048,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_DefEqCache_reducible___default; LEAN_EXPORT lean_object* l_Lean_Meta_setPostponed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_forallBoundedTelescopeImp(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_instMonadLCtxMetaM___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_ParamInfo_isExplicit(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Cache_defEqTrans___default; LEAN_EXPORT lean_object* l_Lean_Meta_ppExprWithInfos(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -34585,125 +34583,6 @@ return x_38; } } } -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; -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; -} -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 -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; 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_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_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: { @@ -34711,76 +34590,269 @@ uint8_t x_11; x_11 = !lean_is_exclusive(x_8); if (x_11 == 0) { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; 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; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; x_12 = lean_ctor_get(x_8, 5); x_13 = l_Lean_replaceRef(x_3, x_12); lean_dec(x_12); lean_ctor_set(x_8, 5, x_13); -x_14 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_4, x_6, x_7, x_8, x_9, x_10); +x_14 = lean_st_ref_get(x_9, x_10); 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___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); +x_17 = lean_ctor_get(x_15, 3); +lean_inc(x_17); +lean_dec(x_15); +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_addTraceNode___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 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_23, x_6, x_7, x_8, x_9, x_16); lean_dec(x_8); -return x_17; +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_st_ref_take(x_9, x_26); +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_27, 1); +lean_inc(x_29); +lean_dec(x_27); +x_30 = !lean_is_exclusive(x_28); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_31 = lean_ctor_get(x_28, 3); +lean_dec(x_31); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_3); +lean_ctor_set(x_32, 1, x_25); +x_33 = l_Lean_PersistentArray_push___rarg(x_1, x_32); +lean_ctor_set(x_28, 3, x_33); +x_34 = lean_st_ref_set(x_9, x_28, x_29); +x_35 = !lean_is_exclusive(x_34); +if (x_35 == 0) +{ +lean_object* x_36; lean_object* x_37; +x_36 = lean_ctor_get(x_34, 0); +lean_dec(x_36); +x_37 = lean_box(0); +lean_ctor_set(x_34, 0, x_37); +return x_34; } 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; uint8_t x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -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); -x_26 = lean_ctor_get(x_8, 8); -x_27 = lean_ctor_get(x_8, 9); -x_28 = lean_ctor_get(x_8, 10); -x_29 = lean_ctor_get_uint8(x_8, sizeof(void*)*11); -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_inc(x_19); -lean_inc(x_18); +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_34, 1); +lean_inc(x_38); +lean_dec(x_34); +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_38); +return x_40; +} +} +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; +x_41 = lean_ctor_get(x_28, 0); +x_42 = lean_ctor_get(x_28, 1); +x_43 = lean_ctor_get(x_28, 2); +x_44 = lean_ctor_get(x_28, 4); +x_45 = lean_ctor_get(x_28, 5); +x_46 = lean_ctor_get(x_28, 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_dec(x_28); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_3); +lean_ctor_set(x_47, 1, x_25); +x_48 = l_Lean_PersistentArray_push___rarg(x_1, x_47); +x_49 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_49, 0, x_41); +lean_ctor_set(x_49, 1, x_42); +lean_ctor_set(x_49, 2, x_43); +lean_ctor_set(x_49, 3, x_48); +lean_ctor_set(x_49, 4, x_44); +lean_ctor_set(x_49, 5, x_45); +lean_ctor_set(x_49, 6, x_46); +x_50 = lean_st_ref_set(x_9, x_49, x_29); +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 = lean_box(0); +if (lean_is_scalar(x_52)) { + x_54 = lean_alloc_ctor(0, 2, 0); +} else { + x_54 = x_52; +} +lean_ctor_set(x_54, 0, x_53); +lean_ctor_set(x_54, 1, x_51); +return x_54; +} +} +else +{ +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; uint8_t x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; size_t x_75; size_t x_76; 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_55 = lean_ctor_get(x_8, 0); +x_56 = lean_ctor_get(x_8, 1); +x_57 = lean_ctor_get(x_8, 2); +x_58 = lean_ctor_get(x_8, 3); +x_59 = lean_ctor_get(x_8, 4); +x_60 = lean_ctor_get(x_8, 5); +x_61 = lean_ctor_get(x_8, 6); +x_62 = lean_ctor_get(x_8, 7); +x_63 = lean_ctor_get(x_8, 8); +x_64 = lean_ctor_get(x_8, 9); +x_65 = lean_ctor_get(x_8, 10); +x_66 = lean_ctor_get_uint8(x_8, sizeof(void*)*11); +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_inc(x_55); lean_dec(x_8); -x_30 = l_Lean_replaceRef(x_3, x_23); -lean_dec(x_23); -x_31 = lean_alloc_ctor(0, 11, 1); -lean_ctor_set(x_31, 0, x_18); -lean_ctor_set(x_31, 1, x_19); -lean_ctor_set(x_31, 2, x_20); -lean_ctor_set(x_31, 3, x_21); -lean_ctor_set(x_31, 4, x_22); -lean_ctor_set(x_31, 5, x_30); -lean_ctor_set(x_31, 6, x_24); -lean_ctor_set(x_31, 7, x_25); -lean_ctor_set(x_31, 8, x_26); -lean_ctor_set(x_31, 9, x_27); -lean_ctor_set(x_31, 10, x_28); -lean_ctor_set_uint8(x_31, sizeof(void*)*11, x_29); -x_32 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_4, x_6, x_7, x_31, x_9, x_10); -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___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_processPostponed___spec__5(x_1, x_2, x_3, x_33, x_5, x_6, x_7, x_31, x_9, x_34); -lean_dec(x_31); -return x_35; +x_67 = l_Lean_replaceRef(x_3, x_60); +lean_dec(x_60); +x_68 = lean_alloc_ctor(0, 11, 1); +lean_ctor_set(x_68, 0, x_55); +lean_ctor_set(x_68, 1, x_56); +lean_ctor_set(x_68, 2, x_57); +lean_ctor_set(x_68, 3, x_58); +lean_ctor_set(x_68, 4, x_59); +lean_ctor_set(x_68, 5, x_67); +lean_ctor_set(x_68, 6, x_61); +lean_ctor_set(x_68, 7, x_62); +lean_ctor_set(x_68, 8, x_63); +lean_ctor_set(x_68, 9, x_64); +lean_ctor_set(x_68, 10, x_65); +lean_ctor_set_uint8(x_68, sizeof(void*)*11, x_66); +x_69 = lean_st_ref_get(x_9, x_10); +x_70 = lean_ctor_get(x_69, 0); +lean_inc(x_70); +x_71 = lean_ctor_get(x_69, 1); +lean_inc(x_71); +lean_dec(x_69); +x_72 = lean_ctor_get(x_70, 3); +lean_inc(x_72); +lean_dec(x_70); +x_73 = l_Lean_PersistentArray_toArray___rarg(x_72); +x_74 = lean_array_get_size(x_73); +x_75 = lean_usize_of_nat(x_74); +lean_dec(x_74); +x_76 = 0; +x_77 = l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNode___spec__1(x_75, x_76, x_73); +x_78 = lean_alloc_ctor(9, 3, 1); +lean_ctor_set(x_78, 0, x_2); +lean_ctor_set(x_78, 1, x_4); +lean_ctor_set(x_78, 2, x_77); +lean_ctor_set_uint8(x_78, sizeof(void*)*3, x_5); +x_79 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_78, x_6, x_7, x_68, x_9, x_71); +lean_dec(x_68); +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 = lean_st_ref_take(x_9, 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_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, 4); +lean_inc(x_88); +x_89 = lean_ctor_get(x_83, 5); +lean_inc(x_89); +x_90 = lean_ctor_get(x_83, 6); +lean_inc(x_90); +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); + lean_ctor_release(x_83, 4); + lean_ctor_release(x_83, 5); + lean_ctor_release(x_83, 6); + x_91 = x_83; +} else { + lean_dec_ref(x_83); + x_91 = lean_box(0); } +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_3); +lean_ctor_set(x_92, 1, x_80); +x_93 = l_Lean_PersistentArray_push___rarg(x_1, x_92); +if (lean_is_scalar(x_91)) { + x_94 = lean_alloc_ctor(0, 7, 0); +} else { + x_94 = x_91; +} +lean_ctor_set(x_94, 0, x_85); +lean_ctor_set(x_94, 1, x_86); +lean_ctor_set(x_94, 2, x_87); +lean_ctor_set(x_94, 3, x_93); +lean_ctor_set(x_94, 4, x_88); +lean_ctor_set(x_94, 5, x_89); +lean_ctor_set(x_94, 6, x_90); +x_95 = lean_st_ref_set(x_9, x_94, x_84); +x_96 = lean_ctor_get(x_95, 1); +lean_inc(x_96); +if (lean_is_exclusive(x_95)) { + lean_ctor_release(x_95, 0); + lean_ctor_release(x_95, 1); + x_97 = x_95; +} else { + lean_dec_ref(x_95); + x_97 = lean_box(0); +} +x_98 = lean_box(0); +if (lean_is_scalar(x_97)) { + x_99 = lean_alloc_ctor(0, 2, 0); +} else { + x_99 = x_97; } +lean_ctor_set(x_99, 0, x_98); +lean_ctor_set(x_99, 1, x_96); +return x_99; } -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) { +} +} +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___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) { _start: { if (lean_obj_tag(x_1) == 0) @@ -34880,7 +34952,7 @@ x_13 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPo 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__6(x_5, x_8, x_9, x_10, x_11, x_14); +x_15 = l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___spec__5(x_5, x_8, x_9, x_10, x_11, x_14); lean_dec(x_10); return x_15; } @@ -35147,7 +35219,7 @@ x_30 = lean_st_ref_set(x_12, x_25, x_26); x_31 = lean_ctor_get(x_30, 1); lean_inc(x_31); lean_dec(x_30); -x_32 = l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___spec__6(x_21, x_9, x_10, x_11, x_12, x_31); +x_32 = l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___spec__5(x_21, x_9, x_10, x_11, x_12, x_31); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -35229,7 +35301,7 @@ x_50 = lean_st_ref_set(x_12, x_49, x_26); x_51 = lean_ctor_get(x_50, 1); lean_inc(x_51); lean_dec(x_50); -x_52 = l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___spec__6(x_21, x_9, x_10, x_11, x_12, x_51); +x_52 = l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___spec__5(x_21, x_9, x_10, x_11, x_12, x_51); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -35609,20 +35681,6 @@ lean_dec(x_1); return x_4; } } -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__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); -lean_dec(x_6); -return x_12; -} -} 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: { @@ -35636,11 +35694,11 @@ lean_dec(x_6); return x_12; } } -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) { +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___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) { _start: { lean_object* x_7; -x_7 = l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___spec__6(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___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); diff --git a/stage0/stdlib/Lean/Meta/ExprDefEq.c b/stage0/stdlib/Lean/Meta/ExprDefEq.c index 17473e38e6fc..df8533261828 100644 --- a/stage0/stdlib/Lean/Meta/ExprDefEq.c +++ b/stage0/stdlib/Lean/Meta/ExprDefEq.c @@ -577,7 +577,6 @@ LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_ExprDe 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_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*); @@ -616,6 +615,7 @@ LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at_Lean_Meta_CheckAssignment_c static lean_object* l_Lean_Meta_CheckAssignment_State_cache___default___closed__1; 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*); +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_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*); @@ -862,7 +862,6 @@ static lean_object* l_Lean_Meta_CheckAssignment_instMonadCacheExprCheckAssignmen 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*); @@ -918,6 +917,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_Ex LEAN_EXPORT lean_object* l_Lean_Meta_DefEqCacheKind_toCtorIdx(uint8_t); lean_object* l_Lean_Meta_setPostponed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDeltaCandidate_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Meta_ParamInfo_isExplicit(lean_object*); static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unstuckMVar___closed__12; LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_Meta_CheckAssignment_checkMVar___spec__15___boxed(lean_object*, lean_object*); @@ -11730,11 +11730,13 @@ 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__5(x_1, x_2, x_3, x_6, x_4, x_8, x_9, x_10, x_11, x_12); +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_processPostponed___spec__6(x_5, x_8, x_9, x_10, x_11, x_14); +x_15 = l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___spec__5(x_5, x_8, x_9, x_10, x_11, x_14); +lean_dec(x_10); return x_15; } } @@ -11801,6 +11803,7 @@ LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_Exp _start: { lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; +lean_dec(x_10); x_16 = l_Lean_exceptBoolEmoji___rarg(x_1); x_17 = l_Lean_stringToMessageData(x_16); lean_dec(x_16); @@ -12034,7 +12037,7 @@ x_49 = lean_st_ref_set(x_11, x_44, x_45); 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_40, x_8, x_9, x_10, x_11, x_50); +x_51 = l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___spec__5(x_40, x_8, x_9, x_10, x_11, x_50); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -12116,7 +12119,7 @@ x_69 = lean_st_ref_set(x_11, x_68, x_45); x_70 = lean_ctor_get(x_69, 1); lean_inc(x_70); lean_dec(x_69); -x_71 = l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___spec__6(x_40, x_8, x_9, x_10, x_11, x_70); +x_71 = l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___spec__5(x_40, x_8, x_9, x_10, x_11, x_70); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -12180,7 +12183,6 @@ x_81 = lean_unbox_float(x_41); lean_dec(x_41); x_82 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3(x_40, x_34, x_14, x_3, x_21, x_4, x_5, x_42, x_81, x_80, x_8, x_9, x_10, x_11, x_39); lean_dec(x_11); -lean_dec(x_10); lean_dec(x_8); return x_82; } @@ -12193,7 +12195,6 @@ x_84 = lean_unbox_float(x_41); lean_dec(x_41); x_85 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3(x_40, x_34, x_14, x_3, x_21, x_4, x_5, x_42, x_84, x_83, x_8, x_9, x_10, x_11, x_39); lean_dec(x_11); -lean_dec(x_10); lean_dec(x_8); return x_85; } @@ -13128,7 +13129,6 @@ x_13 = lean_unbox(x_4); lean_dec(x_4); x_14 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___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_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -13148,9 +13148,7 @@ 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; } } diff --git a/stage0/stdlib/Lean/Meta/Injective.c b/stage0/stdlib/Lean/Meta/Injective.c index f7fc5fc0df10..ac25865ece0c 100644 --- a/stage0/stdlib/Lean/Meta/Injective.c +++ b/stage0/stdlib/Lean/Meta/Injective.c @@ -13,7 +13,7 @@ #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_MonadExcept_ofExcept___at_Lean_Meta_mkInjectiveTheorems___spec__10(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___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*); @@ -24,7 +24,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkAnd_x3f( LEAN_EXPORT lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_throwInjectiveTheoremFailure(lean_object*); extern lean_object* l_Lean_profiler; LEAN_EXPORT lean_object* l_Lean_Meta_elimOptParam___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__12___lambda__1(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_mkArgs2(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___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheoremTypeCore_x3f_mkArgs2___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_mk_empty_array_with_capacity(lean_object*); @@ -34,6 +33,7 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_Injec LEAN_EXPORT lean_object* l_Lean_Meta_genInjectivity; lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_2157____closed__8; +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__11___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentArray_toArray___rarg(lean_object*); lean_object* lean_whnf(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Injective_0__Lean_Meta_solveEqOfCtorEq___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -49,13 +49,11 @@ static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_2157____ double l_Lean_trace_profiler_threshold_getSecs(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*); -static lean_object* l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__12___closed__1; 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___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheoremTypeCore_x3f_mkArgs2___closed__4; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_2157____closed__3; -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__12___lambda__2(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_mkInjectiveTheorems___spec__7___rarg___closed__2; lean_object* l_Lean_replaceRef(lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); @@ -83,10 +81,8 @@ lean_object* l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(lean_ static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1818____closed__5; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_2157____closed__6; LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Meta_Injective_0__Lean_Meta_mkAnd_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_mkInjectiveTheorems___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, 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*); LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1818_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_mkInjectiveTheorems___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_Name_mkStr3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Meta_Injective_0__Lean_Meta_solveEqOfCtorEq___spec__2(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_Injective_0__Lean_Meta_mkInjectiveTheoremTypeCore_x3f___spec__2(size_t, size_t, lean_object*); @@ -104,6 +100,7 @@ lean_object* l_Lean_throwError___at_Lean_Expr_abstractRangeM___spec__1(lean_obje static lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_solveEqOfCtorEq___closed__1; lean_object* lean_st_ref_take(lean_object*, lean_object*); lean_object* l_Lean_Expr_getRevArg_x21(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__11(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__1; LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_mkInjectiveTheorems___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_expr_eqv(lean_object*, lean_object*); @@ -122,7 +119,6 @@ lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean LEAN_EXPORT lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheoremTypeCore_x3f___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveEqTheoremValue___lambda__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_Meta_mkInjectiveTheorems___spec__5___lambda__3___closed__1; -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_mkInjectiveTheorems___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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*); @@ -131,6 +127,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiv lean_object* l_Lean_PersistentArray_append___rarg(lean_object*, lean_object*); 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*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheoremTypeCore_x3f___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__11___closed__1; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_2157____closed__10; static lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveEqTheoremValue___lambda__2___closed__4; static lean_object* l_Lean_Meta_mkInjectiveEqTheoremNameFor___closed__1; @@ -148,9 +145,11 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_mkInjectiveTheorems___ lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); lean_object* l_Lean_Name_append(lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1818____closed__3; +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__11___lambda__1(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___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveEqTheoremValue___lambda__2___closed__2; extern lean_object* l_Lean_instInhabitedExpr; +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__11___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkInjectiveTheorems(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveEqTheoremValue___lambda__2___closed__3; lean_object* l_Lean_MVarId_casesAnd(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -159,6 +158,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiv LEAN_EXPORT lean_object* l_Lean_Meta_elimOptParam___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_2157____closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_2157_(lean_object*); +static lean_object* l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__11___closed__2; lean_object* l_Lean_Meta_splitAndCore(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_Meta_injection(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -177,10 +177,12 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_injTheorem LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__3(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_mkInjectiveTheorems___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_inheritedTraceOptions; lean_object* l_Lean_MessageData_ofExpr(lean_object*); lean_object* l_Lean_MVarId_substEqs(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_mkInjectiveTheorems___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNode___spec__1(size_t, size_t, lean_object*); static lean_object* l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__3___closed__2; static lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_throwInjectiveTheoremFailure___rarg___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_mkInjectiveTheorems___spec__7(lean_object*, lean_object*, lean_object*); @@ -190,7 +192,6 @@ static lean_object* l_Lean_Meta_mkInjectiveEqTheoremNameFor___closed__2; static lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_mkInjectiveTheorems___spec__7___rarg___closed__3; 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*); -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_mkInjectiveTheorems___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_getConstInfo___at_Lean_Meta_mkInjectiveTheorems___spec__2___closed__1; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___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___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveEqTheoremValue(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -213,7 +214,6 @@ LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems 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_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__3___closed__3; lean_object* lean_panic_fn(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__12___lambda__1___boxed(lean_object*, lean_object*, lean_object*, 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_Injective___hyg_1818____closed__6; @@ -242,13 +242,11 @@ static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_2157____ lean_object* lean_array_uget(lean_object*, size_t); extern lean_object* l_Lean_trace_profiler; 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*); -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__12(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_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__4___closed__1; LEAN_EXPORT lean_object* l_Lean_getConstInfoCtor___at_Lean_Meta_mkInjectiveTheorems___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__3___closed__4; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_2157____closed__2; -static lean_object* l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__12___closed__2; lean_object* l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); @@ -4803,125 +4801,6 @@ return x_38; } } } -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_mkInjectiveTheorems___spec__10(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; -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; -} -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 -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; 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_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_Meta_mkInjectiveTheorems___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, lean_object* x_9, lean_object* x_10) { _start: { @@ -4929,76 +4808,269 @@ uint8_t x_11; x_11 = !lean_is_exclusive(x_8); if (x_11 == 0) { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; 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; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; x_12 = lean_ctor_get(x_8, 5); x_13 = l_Lean_replaceRef(x_3, x_12); lean_dec(x_12); lean_ctor_set(x_8, 5, x_13); -x_14 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_4, x_6, x_7, x_8, x_9, x_10); +x_14 = lean_st_ref_get(x_9, x_10); 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___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_mkInjectiveTheorems___spec__10(x_1, x_2, x_3, x_15, x_5, x_6, x_7, x_8, x_9, x_16); +x_17 = lean_ctor_get(x_15, 3); +lean_inc(x_17); +lean_dec(x_15); +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_addTraceNode___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 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_23, x_6, x_7, x_8, x_9, x_16); lean_dec(x_8); -return x_17; +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_st_ref_take(x_9, x_26); +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_27, 1); +lean_inc(x_29); +lean_dec(x_27); +x_30 = !lean_is_exclusive(x_28); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_31 = lean_ctor_get(x_28, 3); +lean_dec(x_31); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_3); +lean_ctor_set(x_32, 1, x_25); +x_33 = l_Lean_PersistentArray_push___rarg(x_1, x_32); +lean_ctor_set(x_28, 3, x_33); +x_34 = lean_st_ref_set(x_9, x_28, x_29); +x_35 = !lean_is_exclusive(x_34); +if (x_35 == 0) +{ +lean_object* x_36; lean_object* x_37; +x_36 = lean_ctor_get(x_34, 0); +lean_dec(x_36); +x_37 = lean_box(0); +lean_ctor_set(x_34, 0, x_37); +return x_34; } 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; uint8_t x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -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); -x_26 = lean_ctor_get(x_8, 8); -x_27 = lean_ctor_get(x_8, 9); -x_28 = lean_ctor_get(x_8, 10); -x_29 = lean_ctor_get_uint8(x_8, sizeof(void*)*11); -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_inc(x_19); -lean_inc(x_18); +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_34, 1); +lean_inc(x_38); +lean_dec(x_34); +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_38); +return x_40; +} +} +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; +x_41 = lean_ctor_get(x_28, 0); +x_42 = lean_ctor_get(x_28, 1); +x_43 = lean_ctor_get(x_28, 2); +x_44 = lean_ctor_get(x_28, 4); +x_45 = lean_ctor_get(x_28, 5); +x_46 = lean_ctor_get(x_28, 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_dec(x_28); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_3); +lean_ctor_set(x_47, 1, x_25); +x_48 = l_Lean_PersistentArray_push___rarg(x_1, x_47); +x_49 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_49, 0, x_41); +lean_ctor_set(x_49, 1, x_42); +lean_ctor_set(x_49, 2, x_43); +lean_ctor_set(x_49, 3, x_48); +lean_ctor_set(x_49, 4, x_44); +lean_ctor_set(x_49, 5, x_45); +lean_ctor_set(x_49, 6, x_46); +x_50 = lean_st_ref_set(x_9, x_49, x_29); +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 = lean_box(0); +if (lean_is_scalar(x_52)) { + x_54 = lean_alloc_ctor(0, 2, 0); +} else { + x_54 = x_52; +} +lean_ctor_set(x_54, 0, x_53); +lean_ctor_set(x_54, 1, x_51); +return x_54; +} +} +else +{ +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; uint8_t x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; size_t x_75; size_t x_76; 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_55 = lean_ctor_get(x_8, 0); +x_56 = lean_ctor_get(x_8, 1); +x_57 = lean_ctor_get(x_8, 2); +x_58 = lean_ctor_get(x_8, 3); +x_59 = lean_ctor_get(x_8, 4); +x_60 = lean_ctor_get(x_8, 5); +x_61 = lean_ctor_get(x_8, 6); +x_62 = lean_ctor_get(x_8, 7); +x_63 = lean_ctor_get(x_8, 8); +x_64 = lean_ctor_get(x_8, 9); +x_65 = lean_ctor_get(x_8, 10); +x_66 = lean_ctor_get_uint8(x_8, sizeof(void*)*11); +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_inc(x_55); lean_dec(x_8); -x_30 = l_Lean_replaceRef(x_3, x_23); -lean_dec(x_23); -x_31 = lean_alloc_ctor(0, 11, 1); -lean_ctor_set(x_31, 0, x_18); -lean_ctor_set(x_31, 1, x_19); -lean_ctor_set(x_31, 2, x_20); -lean_ctor_set(x_31, 3, x_21); -lean_ctor_set(x_31, 4, x_22); -lean_ctor_set(x_31, 5, x_30); -lean_ctor_set(x_31, 6, x_24); -lean_ctor_set(x_31, 7, x_25); -lean_ctor_set(x_31, 8, x_26); -lean_ctor_set(x_31, 9, x_27); -lean_ctor_set(x_31, 10, x_28); -lean_ctor_set_uint8(x_31, sizeof(void*)*11, x_29); -x_32 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_4, x_6, x_7, x_31, x_9, x_10); -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___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_mkInjectiveTheorems___spec__10(x_1, x_2, x_3, x_33, x_5, x_6, x_7, x_31, x_9, x_34); -lean_dec(x_31); -return x_35; +x_67 = l_Lean_replaceRef(x_3, x_60); +lean_dec(x_60); +x_68 = lean_alloc_ctor(0, 11, 1); +lean_ctor_set(x_68, 0, x_55); +lean_ctor_set(x_68, 1, x_56); +lean_ctor_set(x_68, 2, x_57); +lean_ctor_set(x_68, 3, x_58); +lean_ctor_set(x_68, 4, x_59); +lean_ctor_set(x_68, 5, x_67); +lean_ctor_set(x_68, 6, x_61); +lean_ctor_set(x_68, 7, x_62); +lean_ctor_set(x_68, 8, x_63); +lean_ctor_set(x_68, 9, x_64); +lean_ctor_set(x_68, 10, x_65); +lean_ctor_set_uint8(x_68, sizeof(void*)*11, x_66); +x_69 = lean_st_ref_get(x_9, x_10); +x_70 = lean_ctor_get(x_69, 0); +lean_inc(x_70); +x_71 = lean_ctor_get(x_69, 1); +lean_inc(x_71); +lean_dec(x_69); +x_72 = lean_ctor_get(x_70, 3); +lean_inc(x_72); +lean_dec(x_70); +x_73 = l_Lean_PersistentArray_toArray___rarg(x_72); +x_74 = lean_array_get_size(x_73); +x_75 = lean_usize_of_nat(x_74); +lean_dec(x_74); +x_76 = 0; +x_77 = l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNode___spec__1(x_75, x_76, x_73); +x_78 = lean_alloc_ctor(9, 3, 1); +lean_ctor_set(x_78, 0, x_2); +lean_ctor_set(x_78, 1, x_4); +lean_ctor_set(x_78, 2, x_77); +lean_ctor_set_uint8(x_78, sizeof(void*)*3, x_5); +x_79 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_78, x_6, x_7, x_68, x_9, x_71); +lean_dec(x_68); +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 = lean_st_ref_take(x_9, 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_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, 4); +lean_inc(x_88); +x_89 = lean_ctor_get(x_83, 5); +lean_inc(x_89); +x_90 = lean_ctor_get(x_83, 6); +lean_inc(x_90); +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); + lean_ctor_release(x_83, 4); + lean_ctor_release(x_83, 5); + lean_ctor_release(x_83, 6); + x_91 = x_83; +} else { + lean_dec_ref(x_83); + x_91 = lean_box(0); } +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_3); +lean_ctor_set(x_92, 1, x_80); +x_93 = l_Lean_PersistentArray_push___rarg(x_1, x_92); +if (lean_is_scalar(x_91)) { + x_94 = lean_alloc_ctor(0, 7, 0); +} else { + x_94 = x_91; +} +lean_ctor_set(x_94, 0, x_85); +lean_ctor_set(x_94, 1, x_86); +lean_ctor_set(x_94, 2, x_87); +lean_ctor_set(x_94, 3, x_93); +lean_ctor_set(x_94, 4, x_88); +lean_ctor_set(x_94, 5, x_89); +lean_ctor_set(x_94, 6, x_90); +x_95 = lean_st_ref_set(x_9, x_94, x_84); +x_96 = lean_ctor_get(x_95, 1); +lean_inc(x_96); +if (lean_is_exclusive(x_95)) { + lean_ctor_release(x_95, 0); + lean_ctor_release(x_95, 1); + x_97 = x_95; +} else { + lean_dec_ref(x_95); + x_97 = lean_box(0); +} +x_98 = lean_box(0); +if (lean_is_scalar(x_97)) { + x_99 = lean_alloc_ctor(0, 2, 0); +} else { + x_99 = x_97; +} +lean_ctor_set(x_99, 0, x_98); +lean_ctor_set(x_99, 1, x_96); +return x_99; } } -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_mkInjectiveTheorems___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_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_mkInjectiveTheorems___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: { if (lean_obj_tag(x_1) == 0) @@ -5098,7 +5170,7 @@ x_13 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_mkInjecti x_14 = lean_ctor_get(x_13, 1); lean_inc(x_14); lean_dec(x_13); -x_15 = l_MonadExcept_ofExcept___at_Lean_Meta_mkInjectiveTheorems___spec__11(x_5, x_8, x_9, x_10, x_11, x_14); +x_15 = l_MonadExcept_ofExcept___at_Lean_Meta_mkInjectiveTheorems___spec__10(x_5, x_8, x_9, x_10, x_11, x_14); lean_dec(x_10); return x_15; } @@ -5365,7 +5437,7 @@ x_30 = lean_st_ref_set(x_12, x_25, x_26); x_31 = lean_ctor_get(x_30, 1); lean_inc(x_31); lean_dec(x_30); -x_32 = l_MonadExcept_ofExcept___at_Lean_Meta_mkInjectiveTheorems___spec__11(x_21, x_9, x_10, x_11, x_12, x_31); +x_32 = l_MonadExcept_ofExcept___at_Lean_Meta_mkInjectiveTheorems___spec__10(x_21, x_9, x_10, x_11, x_12, x_31); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -5447,7 +5519,7 @@ x_50 = lean_st_ref_set(x_12, x_49, x_26); x_51 = lean_ctor_get(x_50, 1); lean_inc(x_51); lean_dec(x_50); -x_52 = l_MonadExcept_ofExcept___at_Lean_Meta_mkInjectiveTheorems___spec__11(x_21, x_9, x_10, x_11, x_12, x_51); +x_52 = l_MonadExcept_ofExcept___at_Lean_Meta_mkInjectiveTheorems___spec__10(x_21, x_9, x_10, x_11, x_12, x_51); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -5652,7 +5724,7 @@ return x_32; } } } -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__12___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_Meta_mkInjectiveTheorems___spec__11___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; lean_object* x_12; @@ -5670,7 +5742,7 @@ lean_ctor_set(x_12, 1, x_7); return x_12; } } -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__12___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_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__11___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; @@ -5854,7 +5926,7 @@ return x_39; } } } -static lean_object* _init_l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__12___closed__1() { +static lean_object* _init_l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__11___closed__1() { _start: { lean_object* x_1; @@ -5862,17 +5934,17 @@ x_1 = lean_mk_string_from_bytes("injective", 9); return x_1; } } -static lean_object* _init_l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__12___closed__2() { +static lean_object* _init_l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__11___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1818____closed__6; -x_2 = l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__12___closed__1; +x_2 = l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__11___closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___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_EXPORT lean_object* l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___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) { _start: { if (lean_obj_tag(x_1) == 0) @@ -5897,11 +5969,11 @@ x_10 = lean_ctor_get(x_1, 1); lean_inc(x_10); lean_dec(x_1); lean_inc(x_9); -x_11 = lean_alloc_closure((void*)(l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__12___lambda__1___boxed), 7, 1); +x_11 = lean_alloc_closure((void*)(l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__11___lambda__1___boxed), 7, 1); lean_closure_set(x_11, 0, x_9); -x_12 = lean_alloc_closure((void*)(l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__12___lambda__2), 6, 1); +x_12 = lean_alloc_closure((void*)(l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__11___lambda__2), 6, 1); lean_closure_set(x_12, 0, x_9); -x_13 = l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__12___closed__2; +x_13 = l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__11___closed__2; x_14 = 1; lean_inc(x_6); lean_inc(x_5); @@ -5955,7 +6027,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_mkInjectiveTheorems___lambda__1(lean_object { lean_object* x_7; lean_object* x_8; x_7 = lean_box(0); -x_8 = l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__12(x_1, x_7, x_2, x_3, x_4, x_5, x_6); +x_8 = l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__11(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; @@ -6490,20 +6562,6 @@ lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_mkInjectiveTheorems___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: -{ -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_mkInjectiveTheorems___spec__10(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_Meta_mkInjectiveTheorems___spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { @@ -6517,11 +6575,11 @@ lean_dec(x_6); return x_12; } } -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_mkInjectiveTheorems___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_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_mkInjectiveTheorems___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) { _start: { lean_object* x_7; -x_7 = l_MonadExcept_ofExcept___at_Lean_Meta_mkInjectiveTheorems___spec__11(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_MonadExcept_ofExcept___at_Lean_Meta_mkInjectiveTheorems___spec__10(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); @@ -6598,11 +6656,11 @@ x_11 = l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5(x_1, x_ return x_11; } } -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___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_EXPORT lean_object* l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___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, lean_object* x_6, lean_object* x_7) { _start: { lean_object* x_8; -x_8 = l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__12___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +x_8 = l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__11___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); @@ -6737,7 +6795,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_215 _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_Meta_mkInjectiveTheorems___spec__12___closed__2; +x_2 = l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__11___closed__2; x_3 = 0; x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_2157____closed__13; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); @@ -6905,10 +6963,10 @@ l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__4___ lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__4___closed__2); l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__5___closed__1 = _init_l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__5___closed__1(); lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__5___closed__1); -l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__12___closed__1 = _init_l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__12___closed__1(); -lean_mark_persistent(l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__12___closed__1); -l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__12___closed__2 = _init_l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__12___closed__2(); -lean_mark_persistent(l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__12___closed__2); +l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__11___closed__1 = _init_l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__11___closed__1(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__11___closed__1); +l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__11___closed__2 = _init_l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__11___closed__2(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__11___closed__2); 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(); diff --git a/stage0/stdlib/Lean/Meta/SynthInstance.c b/stage0/stdlib/Lean/Meta/SynthInstance.c index ed4b2f784c27..9cc990bed5fd 100644 --- a/stage0/stdlib/Lean/Meta/SynthInstance.c +++ b/stage0/stdlib/Lean/Meta/SynthInstance.c @@ -15,13 +15,13 @@ extern "C" { #endif extern lean_object* l_Lean_Meta_isDefEqStuckExceptionId; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_SynthInstance_consume___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* 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_Meta_SynthInstance_consume___spec__3(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_synthInstance_x3f___spec__2___boxed(lean_object*, lean_object*, lean_object*); 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_9390____closed__12; 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*); @@ -62,6 +62,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_consume___lambda__3___boxed(l 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_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_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__15(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*); static lean_object* l_Lean_Meta_synthInstance_x3f___lambda__6___closed__8; @@ -69,7 +70,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_generate___lambda__4(lean_obj LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_resume___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_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___spec__3(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_addAnswer___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*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__9; 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*); @@ -87,7 +87,6 @@ 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; @@ -95,7 +94,6 @@ static lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal LEAN_EXPORT lean_object* l_StateT_bind___at_Lean_Meta_SynthInstance_MkTableKey_instMonadMCtxM___spec__2___rarg(lean_object*, lean_object*, lean_object*); 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_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*); @@ -134,10 +132,10 @@ 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_AssocList_foldlM___at_Lean_Meta_SynthInstance_newSubgoal___spec__5(lean_object*, lean_object*); lean_object* l_Lean_replaceRef(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__11; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, 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* l_Lean_Expr_proj___override(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_bvar___override(lean_object*); @@ -194,6 +192,7 @@ static lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendin uint8_t l___private_Lean_Util_Trace_0__Lean_checkTraceOption(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessLevels___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_synthInstance_x3f_assignOutParams___closed__1; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__10; 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*); @@ -201,19 +200,20 @@ 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*); extern lean_object* l_Lean_Expr_instBEqExpr; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__3; 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_9390____closed__8; 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*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___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___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessArgs___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_Meta_SynthInstance_getInstances___lambda__1(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_SynthInstance_mkLambdaFVars_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_synthInstance_x3f_assignOutParams___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_SynthInstance_wakeUp___closed__2; static lean_object* l_Lean_Meta_synthInstance_x3f___lambda__6___closed__1; @@ -221,6 +221,7 @@ 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; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__5; 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_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; @@ -236,6 +237,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_MonadExcept_ofExcept___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_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*, 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*); @@ -329,6 +331,7 @@ 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_9390____closed__1; 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*); @@ -336,7 +339,6 @@ LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_addAn lean_object* l_Lean_Meta_abstractMVars(lean_object*, uint8_t, 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*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_consume(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMapImp_moveEntries___at_Lean_Meta_SynthInstance_newSubgoal___spec__4(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); @@ -372,8 +374,8 @@ LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_getNextToResume___rarg(lean_o 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; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__2; lean_object* l_Lean_Expr_constName_x3f(lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__12; 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*); @@ -387,6 +389,7 @@ 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_newSubgoal___spec__14___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_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*); @@ -395,7 +398,6 @@ lean_object* l_Lean_Expr_letE___override(lean_object*, lean_object*, lean_object 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*); uint8_t lean_name_eq(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__1; 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*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_mkTableKeyFor(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -410,6 +412,7 @@ static lean_object* l_Lean_Meta_SynthInstance_instInhabitedInstance___closed__3; 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_9390____closed__9; 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*); @@ -426,13 +429,13 @@ 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_Meta_SynthInstance_instInhabitedSynthM___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -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*); 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_SynthInstance_mkLambdaFVars_x27___boxed(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_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*); LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -452,14 +455,12 @@ LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_generate(lean_object*, lean_o 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; static lean_object* l_Lean_Meta_SynthInstance_instInhabitedSynthM___rarg___closed__1; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_synthInstance_x3f___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_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; 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_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____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*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_arrayExpr_toMessageData(lean_object*, lean_object*, lean_object*); @@ -473,6 +474,7 @@ LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_tryRe 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_hashmap_mk_idx(lean_object*, uint64_t); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__3; LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_Meta_SynthInstance_findEntry_x3f___spec__2(lean_object*, lean_object*); lean_object* l_Lean_Meta_instInhabitedMetaM___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_findEntry_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -487,8 +489,8 @@ 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_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*); +lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNode___spec__1(size_t, size_t, lean_object*); 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*); @@ -496,12 +498,12 @@ static lean_object* l_Lean_Meta_SynthInstance_instInhabitedGeneratorNode___close 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*); LEAN_EXPORT lean_object* l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_tryResolve___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_eta(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_findEntry_x3f___boxed(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*); 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_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__13; static lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__2___closed__3; 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*); @@ -518,8 +520,11 @@ uint8_t l_Lean_Expr_hasAnyFVar_visit___at_Lean_Expr_containsFVar___spec__1(lean_ 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*); 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_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__6; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__13; static lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__2___closed__2; static lean_object* l_Lean_Meta_SynthInstance_checkSystem___closed__4; +LEAN_EXPORT lean_object* l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__15___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_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*); 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*); @@ -533,12 +538,10 @@ LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_tryResolve___lambda__4(lean_o static lean_object* l_Lean_Meta_SynthInstance_MkTableKey_State_lmap___default___closed__1; 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_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__10; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_44____closed__1; LEAN_EXPORT lean_object* l_Lean_HashMap_insert___at_Lean_Meta_SynthInstance_newSubgoal___spec__1(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_getFVarLocalDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_Meta_SynthInstance_findEntry_x3f___spec__2___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353_(lean_object*); uint8_t lean_nat_dec_lt(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*); @@ -562,7 +565,6 @@ uint8_t l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_399_( 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___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; @@ -607,7 +609,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_main___lambda__1(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*); static lean_object* l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__1; static lean_object* l_Lean_Meta_SynthInstance_tryResolve___lambda__5___closed__1; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__5; 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*); 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*); @@ -632,7 +633,6 @@ static lean_object* l_Lean_Meta_SynthInstance_generate___lambda__1___closed__2; lean_object* l_Lean_Meta_isClass_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkHashMapImp___rarg(lean_object*); lean_object* l_Lean_instantiateMVarsCore(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_synthInstance_x3f_assignOutParams___lambda__1(uint8_t, 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_main___spec__1(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___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -685,6 +685,7 @@ uint8_t l_Lean_PersistentHashMap_contains___at_Lean_NameSSet_contains___spec__3( 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*, uint8_t, lean_object*, lean_object*, uint8_t, double, 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_addAnswer___lambda__4___closed__4; +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390_(lean_object*); 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*); static lean_object* l_Lean_Meta_SynthInstance_checkSystem___closed__5; lean_object* lean_array_get_size(lean_object*); @@ -707,15 +708,16 @@ LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Meta_synthInstance_x3f___spec__5( 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; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____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_Meta_SynthInstance_getInstances___lambda__2___closed__3; lean_object* lean_nat_add(lean_object*, lean_object*); static lean_object* l_Lean_Meta_SynthInstance_main___lambda__1___closed__3; lean_object* l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__2; static lean_object* l_Lean_withTraceNode_x27___at_Lean_Meta_SynthInstance_newSubgoal___spec__8___closed__1; uint8_t l_Lean_Exception_isRuntime(lean_object*); lean_object* l_Lean_Expr_lam___override(lean_object*, lean_object*, lean_object*, uint8_t); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__4; static lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___lambda__2___closed__1; lean_object* l_Array_instBEqArray___rarg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_SynthInstance_instInhabitedGeneratorNode___closed__2; @@ -758,7 +760,6 @@ 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; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__11; 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*); @@ -5794,125 +5795,6 @@ return x_40; } } } -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; -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_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_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_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; -} -} -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_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_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: { @@ -5920,76 +5802,269 @@ 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; +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; size_t x_22; size_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; 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_16 = lean_st_ref_get(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_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); +x_19 = lean_ctor_get(x_17, 3); +lean_inc(x_19); +lean_dec(x_17); +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_addTraceNode___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 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_25, x_8, x_9, x_10, x_11, x_18); lean_dec(x_10); -return x_19; +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_st_ref_take(x_11, 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_is_exclusive(x_30); +if (x_32 == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_33 = lean_ctor_get(x_30, 3); +lean_dec(x_33); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_3); +lean_ctor_set(x_34, 1, x_27); +x_35 = l_Lean_PersistentArray_push___rarg(x_1, x_34); +lean_ctor_set(x_30, 3, x_35); +x_36 = lean_st_ref_set(x_11, x_30, x_31); +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 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; } 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; uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -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); -x_31 = lean_ctor_get_uint8(x_10, sizeof(void*)*11); -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_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; +} +} +else +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_43 = lean_ctor_get(x_30, 0); +x_44 = lean_ctor_get(x_30, 1); +x_45 = lean_ctor_get(x_30, 2); +x_46 = lean_ctor_get(x_30, 4); +x_47 = lean_ctor_get(x_30, 5); +x_48 = lean_ctor_get(x_30, 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_dec(x_30); +x_49 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_49, 0, x_3); +lean_ctor_set(x_49, 1, x_27); +x_50 = l_Lean_PersistentArray_push___rarg(x_1, x_49); +x_51 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_51, 0, x_43); +lean_ctor_set(x_51, 1, x_44); +lean_ctor_set(x_51, 2, x_45); +lean_ctor_set(x_51, 3, x_50); +lean_ctor_set(x_51, 4, x_46); +lean_ctor_set(x_51, 5, x_47); +lean_ctor_set(x_51, 6, x_48); +x_52 = lean_st_ref_set(x_11, x_51, x_31); +x_53 = lean_ctor_get(x_52, 1); +lean_inc(x_53); +if (lean_is_exclusive(x_52)) { + lean_ctor_release(x_52, 0); + lean_ctor_release(x_52, 1); + x_54 = x_52; +} else { + lean_dec_ref(x_52); + x_54 = lean_box(0); +} +x_55 = lean_box(0); +if (lean_is_scalar(x_54)) { + x_56 = lean_alloc_ctor(0, 2, 0); +} else { + x_56 = x_54; +} +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_53); +return x_56; +} +} +else +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; uint8_t x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; size_t x_77; size_t x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; +x_57 = lean_ctor_get(x_10, 0); +x_58 = lean_ctor_get(x_10, 1); +x_59 = lean_ctor_get(x_10, 2); +x_60 = lean_ctor_get(x_10, 3); +x_61 = lean_ctor_get(x_10, 4); +x_62 = lean_ctor_get(x_10, 5); +x_63 = lean_ctor_get(x_10, 6); +x_64 = lean_ctor_get(x_10, 7); +x_65 = lean_ctor_get(x_10, 8); +x_66 = lean_ctor_get(x_10, 9); +x_67 = lean_ctor_get(x_10, 10); +x_68 = lean_ctor_get_uint8(x_10, sizeof(void*)*11); +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_dec(x_10); -x_32 = l_Lean_replaceRef(x_3, x_25); -lean_dec(x_25); -x_33 = lean_alloc_ctor(0, 11, 1); -lean_ctor_set(x_33, 0, x_20); -lean_ctor_set(x_33, 1, x_21); -lean_ctor_set(x_33, 2, x_22); -lean_ctor_set(x_33, 3, x_23); -lean_ctor_set(x_33, 4, x_24); -lean_ctor_set(x_33, 5, x_32); -lean_ctor_set(x_33, 6, x_26); -lean_ctor_set(x_33, 7, x_27); -lean_ctor_set(x_33, 8, x_28); -lean_ctor_set(x_33, 9, x_29); -lean_ctor_set(x_33, 10, x_30); -lean_ctor_set_uint8(x_33, sizeof(void*)*11, x_31); -x_34 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_4, x_8, x_9, x_33, x_11, x_12); -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___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_SynthInstance_newSubgoal___spec__14(x_1, x_2, x_3, x_35, x_5, x_6, x_7, x_8, x_9, x_33, x_11, x_36); -lean_dec(x_33); -return x_37; +x_69 = l_Lean_replaceRef(x_3, x_62); +lean_dec(x_62); +x_70 = lean_alloc_ctor(0, 11, 1); +lean_ctor_set(x_70, 0, x_57); +lean_ctor_set(x_70, 1, x_58); +lean_ctor_set(x_70, 2, x_59); +lean_ctor_set(x_70, 3, x_60); +lean_ctor_set(x_70, 4, x_61); +lean_ctor_set(x_70, 5, x_69); +lean_ctor_set(x_70, 6, x_63); +lean_ctor_set(x_70, 7, x_64); +lean_ctor_set(x_70, 8, x_65); +lean_ctor_set(x_70, 9, x_66); +lean_ctor_set(x_70, 10, x_67); +lean_ctor_set_uint8(x_70, sizeof(void*)*11, x_68); +x_71 = lean_st_ref_get(x_11, x_12); +x_72 = lean_ctor_get(x_71, 0); +lean_inc(x_72); +x_73 = lean_ctor_get(x_71, 1); +lean_inc(x_73); +lean_dec(x_71); +x_74 = lean_ctor_get(x_72, 3); +lean_inc(x_74); +lean_dec(x_72); +x_75 = l_Lean_PersistentArray_toArray___rarg(x_74); +x_76 = lean_array_get_size(x_75); +x_77 = lean_usize_of_nat(x_76); +lean_dec(x_76); +x_78 = 0; +x_79 = l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNode___spec__1(x_77, x_78, x_75); +x_80 = lean_alloc_ctor(9, 3, 1); +lean_ctor_set(x_80, 0, x_2); +lean_ctor_set(x_80, 1, x_4); +lean_ctor_set(x_80, 2, x_79); +lean_ctor_set_uint8(x_80, sizeof(void*)*3, x_5); +x_81 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_80, x_8, x_9, x_70, x_11, x_73); +lean_dec(x_70); +x_82 = lean_ctor_get(x_81, 0); +lean_inc(x_82); +x_83 = lean_ctor_get(x_81, 1); +lean_inc(x_83); +lean_dec(x_81); +x_84 = lean_st_ref_take(x_11, x_83); +x_85 = lean_ctor_get(x_84, 0); +lean_inc(x_85); +x_86 = lean_ctor_get(x_84, 1); +lean_inc(x_86); +lean_dec(x_84); +x_87 = lean_ctor_get(x_85, 0); +lean_inc(x_87); +x_88 = lean_ctor_get(x_85, 1); +lean_inc(x_88); +x_89 = lean_ctor_get(x_85, 2); +lean_inc(x_89); +x_90 = lean_ctor_get(x_85, 4); +lean_inc(x_90); +x_91 = lean_ctor_get(x_85, 5); +lean_inc(x_91); +x_92 = lean_ctor_get(x_85, 6); +lean_inc(x_92); +if (lean_is_exclusive(x_85)) { + lean_ctor_release(x_85, 0); + lean_ctor_release(x_85, 1); + lean_ctor_release(x_85, 2); + lean_ctor_release(x_85, 3); + lean_ctor_release(x_85, 4); + lean_ctor_release(x_85, 5); + lean_ctor_release(x_85, 6); + x_93 = x_85; +} else { + lean_dec_ref(x_85); + x_93 = lean_box(0); +} +x_94 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_94, 0, x_3); +lean_ctor_set(x_94, 1, x_82); +x_95 = l_Lean_PersistentArray_push___rarg(x_1, x_94); +if (lean_is_scalar(x_93)) { + x_96 = lean_alloc_ctor(0, 7, 0); +} else { + x_96 = x_93; +} +lean_ctor_set(x_96, 0, x_87); +lean_ctor_set(x_96, 1, x_88); +lean_ctor_set(x_96, 2, x_89); +lean_ctor_set(x_96, 3, x_95); +lean_ctor_set(x_96, 4, x_90); +lean_ctor_set(x_96, 5, x_91); +lean_ctor_set(x_96, 6, x_92); +x_97 = lean_st_ref_set(x_11, x_96, x_86); +x_98 = lean_ctor_get(x_97, 1); +lean_inc(x_98); +if (lean_is_exclusive(x_97)) { + lean_ctor_release(x_97, 0); + lean_ctor_release(x_97, 1); + x_99 = x_97; +} else { + lean_dec_ref(x_97); + x_99 = lean_box(0); +} +x_100 = lean_box(0); +if (lean_is_scalar(x_99)) { + x_101 = lean_alloc_ctor(0, 2, 0); +} else { + x_101 = x_99; +} +lean_ctor_set(x_101, 0, x_100); +lean_ctor_set(x_101, 1, x_98); +return x_101; } } } -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) { +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___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) { _start: { if (lean_obj_tag(x_1) == 0) @@ -6089,7 +6164,7 @@ x_15 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_SynthInst 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__15(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__14(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; @@ -6364,7 +6439,7 @@ x_32 = lean_st_ref_set(x_14, x_27, x_28); x_33 = lean_ctor_get(x_32, 1); lean_inc(x_33); lean_dec(x_32); -x_34 = l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_newSubgoal___spec__15(x_23, x_9, x_10, x_11, x_12, x_13, x_14, x_33); +x_34 = l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_newSubgoal___spec__14(x_23, x_9, x_10, x_11, x_12, x_13, x_14, x_33); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); @@ -6448,7 +6523,7 @@ x_52 = lean_st_ref_set(x_14, x_51, x_28); x_53 = lean_ctor_get(x_52, 1); lean_inc(x_53); lean_dec(x_52); -x_54 = l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_newSubgoal___spec__15(x_23, x_9, x_10, x_11, x_12, x_13, x_14, x_53); +x_54 = l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_newSubgoal___spec__14(x_23, x_9, x_10, x_11, x_12, x_13, x_14, x_53); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); @@ -6757,7 +6832,7 @@ 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) { +LEAN_EXPORT lean_object* l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__15___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; @@ -6809,11 +6884,11 @@ return x_19; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__16(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__15(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); +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__15___rarg), 9, 0); return x_2; } } @@ -7057,7 +7132,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__16___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__15___rarg(x_1, x_18, x_5, x_6, x_7, x_8, x_9, x_10, x_11); return x_19; } } @@ -7108,22 +7183,6 @@ lean_dec(x_1); return x_6; } } -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__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); -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_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: { @@ -7139,11 +7198,11 @@ lean_dec(x_6); return x_14; } } -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) { +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___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) { _start: { lean_object* x_9; -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); +x_9 = l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_newSubgoal___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); @@ -7763,7 +7822,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__16___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__15___rarg(x_1, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_13; } } @@ -8312,6 +8371,77 @@ lean_dec(x_1); return x_7; } } +LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkLambdaFVars_x27(lean_object* x_1, 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; uint8_t x_9; uint8_t x_10; lean_object* x_11; +x_8 = 0; +x_9 = 1; +x_10 = 1; +x_11 = l_Lean_Meta_mkLambdaFVars(x_1, x_2, x_8, x_9, x_10, x_3, x_4, x_5, x_6, x_7); +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 = l_Lean_Expr_eta(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 = l_Lean_Expr_eta(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) +{ +return x_11; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_11, 0); +x_21 = lean_ctor_get(x_11, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_11); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +return x_22; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkLambdaFVars_x27___boxed(lean_object* x_1, 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_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkLambdaFVars_x27(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_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_tryResolve___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) { _start: { @@ -9235,153 +9365,150 @@ return x_18; } 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_19; +x_19 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkLambdaFVars_x27(x_3, x_4, x_7, x_8, x_9, x_10, x_14); +if (lean_obj_tag(x_19) == 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_object* x_20; lean_object* x_21; lean_object* 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); 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) +x_22 = l_Lean_Meta_isExprDefEq(x_5, x_20, x_7, x_8, x_9, x_10, x_21); +if (lean_obj_tag(x_22) == 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_23; uint8_t x_24; +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_unbox(x_23); +lean_dec(x_23); +if (x_24 == 0) { -lean_object* x_28; lean_object* x_29; lean_object* x_30; +lean_object* x_25; lean_object* x_26; lean_object* x_27; 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; +x_25 = lean_ctor_get(x_22, 1); +lean_inc(x_25); +lean_dec(x_22); +x_26 = lean_box(0); +x_27 = lean_apply_6(x_15, x_26, x_7, x_8, x_9, x_10, x_25); +return x_27; } else { -lean_object* x_31; lean_object* x_32; uint8_t x_33; +lean_object* x_28; lean_object* x_29; uint8_t x_30; 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) +x_28 = lean_ctor_get(x_22, 1); +lean_inc(x_28); +lean_dec(x_22); +x_29 = lean_st_ref_get(x_8, x_28); +lean_dec(x_8); +x_30 = !lean_is_exclusive(x_29); +if (x_30 == 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_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_31 = lean_ctor_get(x_29, 0); +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +lean_dec(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_6); +x_34 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_29, 0, x_34); +return x_29; } 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_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +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_ctor_get(x_35, 0); +lean_inc(x_37); +lean_dec(x_35); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_6); +x_39 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_39, 0, x_38); +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_36); +return x_40; } } } else { -uint8_t x_44; +uint8_t x_41; 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) +x_41 = !lean_is_exclusive(x_22); +if (x_41 == 0) { -return x_25; +return x_22; } 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_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_22, 0); +x_43 = lean_ctor_get(x_22, 1); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_22); +x_44 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_44, 0, x_42); +lean_ctor_set(x_44, 1, x_43); +return x_44; } } } else { -uint8_t x_48; +uint8_t x_45; 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) +x_45 = !lean_is_exclusive(x_19); +if (x_45 == 0) { -return x_22; +return x_19; } 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_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_ctor_get(x_19, 0); +x_47 = lean_ctor_get(x_19, 1); +lean_inc(x_47); +lean_inc(x_46); +lean_dec(x_19); +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 { -uint8_t x_52; +uint8_t x_49; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -9390,23 +9517,23 @@ 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) +x_49 = !lean_is_exclusive(x_12); +if (x_49 == 0) { return x_12; } 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_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_12, 0); +x_51 = lean_ctor_get(x_12, 1); +lean_inc(x_51); +lean_inc(x_50); 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_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; } } } @@ -9818,7 +9945,7 @@ 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); +x_15 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__15___rarg(x_1, x_14, x_4, x_5, x_6, x_7, x_8, x_9, x_10); return x_15; } } @@ -12383,7 +12510,7 @@ 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); +x_13 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__15___rarg(x_9, x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8); return x_13; } } @@ -13075,148 +13202,145 @@ 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) { _start: { -lean_object* x_11; uint8_t x_12; uint8_t x_13; uint8_t x_14; lean_object* x_15; +lean_object* x_11; lean_object* x_12; 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) +x_12 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkLambdaFVars_x27(x_2, x_11, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(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; +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_newSubgoal___lambda__2___closed__3; +x_16 = lean_array_push(x_15, x_5); +x_17 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkLambdaFVars_x27(x_16, x_13, x_6, x_7, x_8, x_9, x_14); +if (lean_obj_tag(x_17) == 0) { -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_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__2___closed__2; +x_21 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_20, x_6, x_7, x_8, x_9, x_19); +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); -lean_dec(x_25); -if (x_26 == 0) +x_23 = lean_unbox(x_22); +lean_dec(x_22); +if (x_23 == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; +lean_object* x_24; lean_object* x_25; lean_object* x_26; 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); +x_24 = lean_ctor_get(x_21, 1); +lean_inc(x_24); +lean_dec(x_21); +x_25 = lean_box(0); +x_26 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__1(x_3, x_18, x_25, x_6, x_7, x_8, x_9, x_24); lean_dec(x_8); lean_dec(x_6); -return x_29; +return x_26; } 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_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_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; 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_27 = lean_ctor_get(x_21, 1); +lean_inc(x_27); +lean_dec(x_21); +x_28 = l_Lean_MessageData_ofExpr(x_4); +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_29); +lean_ctor_set(x_30, 1, x_28); +x_31 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__2___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); 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_33 = l_Lean_indentExpr(x_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 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__2___closed__6; +x_36 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +lean_inc(x_18); +x_37 = l_Lean_indentExpr(x_18); +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_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_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_29); +x_40 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_20, x_39, x_6, x_7, x_8, x_9, x_27); +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___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__1(x_3, x_18, x_41, x_6, x_7, x_8, x_9, x_42); lean_dec(x_8); lean_dec(x_6); -lean_dec(x_44); -return x_46; +lean_dec(x_41); +return x_43; } } else { -uint8_t x_47; +uint8_t x_44; 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) +x_44 = !lean_is_exclusive(x_17); +if (x_44 == 0) { -return x_20; +return x_17; } 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_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_17, 0); +x_46 = lean_ctor_get(x_17, 1); +lean_inc(x_46); +lean_inc(x_45); +lean_dec(x_17); +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_51; +uint8_t x_48; 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) +x_48 = !lean_is_exclusive(x_12); +if (x_48 == 0) { -return x_15; +return x_12; } else { -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_49; lean_object* x_50; lean_object* x_51; +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_dec(x_12); +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; } } } @@ -13952,7 +14076,7 @@ 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); +x_23 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__15___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; @@ -14162,7 +14286,7 @@ 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); +x_19 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__15___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; @@ -14276,7 +14400,7 @@ 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); +x_41 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__15___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; @@ -14311,7 +14435,7 @@ 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); +x_50 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__15___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; @@ -14957,7 +15081,7 @@ 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); +x_199 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__15___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; @@ -14993,7 +15117,7 @@ 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); +x_209 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__15___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; @@ -15880,7 +16004,7 @@ 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); +x_402 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__15___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; @@ -16003,7 +16127,7 @@ 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); +x_425 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__15___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; @@ -16043,7 +16167,7 @@ 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); +x_435 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__15___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; @@ -18100,7 +18224,7 @@ 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); +x_31 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__15___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; @@ -18690,7 +18814,7 @@ 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); +x_17 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__15___rarg(x_4, x_16, x_6, x_7, x_8, x_9, x_10, x_11, x_12); return x_17; } } @@ -18880,7 +19004,7 @@ static lean_object* _init_l_Lean_Meta_SynthInstance_resume___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_resume___closed__1; -x_3 = lean_unsigned_to_nat(542u); +x_3 = lean_unsigned_to_nat(562u); 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); @@ -19012,7 +19136,7 @@ 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); +x_36 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__15___rarg(x_32, x_35, x_1, x_2, x_3, x_4, x_5, x_6, x_31); return x_36; } } @@ -27453,7 +27577,7 @@ lean_dec(x_3); return x_9; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__1() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -27463,17 +27587,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_9353____closed__2() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____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_9353____closed__1; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____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_9353____closed__3() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__3() { _start: { lean_object* x_1; @@ -27481,17 +27605,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_9353____closed__4() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____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_9353____closed__2; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__3; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__2; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____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_9353____closed__5() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__5() { _start: { lean_object* x_1; @@ -27499,37 +27623,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_9353____closed__6() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____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_9353____closed__4; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__5; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__4; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____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_9353____closed__7() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____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_9353____closed__6; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____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_9353____closed__8() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____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_9353____closed__7; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____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_9353____closed__9() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__9() { _start: { lean_object* x_1; @@ -27537,17 +27661,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_9353____closed__10() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____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_9353____closed__8; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__9; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__8; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____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_9353____closed__11() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__11() { _start: { lean_object* x_1; @@ -27555,33 +27679,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_9353____closed__12() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____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_9353____closed__10; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__11; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__10; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____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_9353____closed__13() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____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_9353____closed__12; -x_2 = lean_unsigned_to_nat(9353u); +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__12; +x_2 = lean_unsigned_to_nat(9390u); 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_9353_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390_(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_9353____closed__13; +x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__13; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); if (lean_obj_tag(x_5) == 0) { @@ -28205,33 +28329,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_9353____closed__1 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__1(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__1); -l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__2 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__2(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__2); -l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__3 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__3(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__3); -l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__4 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__4(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__4); -l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__5 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__5(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__5); -l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__6 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__6(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__6); -l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__7 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__7(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__7); -l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__8 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__8(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__8); -l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__9 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__9(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__9); -l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__10 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__10(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__10); -l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__11 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__11(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__11); -l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__12 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__12(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__12); -l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__13 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__13(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353____closed__13); -if (builtin) {res = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9353_(lean_io_mk_world()); +l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__1 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__1(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__1); +l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__2 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__2(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__2); +l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__3 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__3(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__3); +l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__4 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__4(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__4); +l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__5 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__5(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__5); +l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__6 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__6(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__6); +l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__7 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__7(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__7); +l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__8 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__8(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__8); +l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__9 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__9(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__9); +l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__10 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__10(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__10); +l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__11 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__11(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__11); +l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__12 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__12(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__12); +l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__13 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__13(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390____closed__13); +if (builtin) {res = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9390_(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/Simp/Rewrite.c b/stage0/stdlib/Lean/Meta/Tactic/Simp/Rewrite.c index f411ad77278d..5941054f8873 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Simp/Rewrite.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Simp/Rewrite.c @@ -14,7 +14,6 @@ extern "C" { #endif LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_tryTheoremCore_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_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*); LEAN_EXPORT lean_object* l_Array_findSomeRevM_x3f_find___at_Lean_Meta_Simp_dischargeUsingAssumption_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* l_Lean_Meta_Simp_mkCongrFun(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_mkSEvalMethods(lean_object*); @@ -68,7 +67,6 @@ lean_object* l_Lean_Meta_Simp_getDtConfig(lean_object*); lean_object* l_Lean_PersistentArray_toArray___rarg(lean_object*); static lean_object* l_Lean_Meta_Simp_synthesizeArgs_synthesizeInstance___closed__9; lean_object* l_Lean_Meta_unifyEq_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_MonadExcept_ofExcept___at_Lean_Meta_Simp_discharge_x3f_x27___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_Array_forInUnsafe_loop___at_Lean_Meta_Simp_sevalGround___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_whnf(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_panic___at___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_tryTheoremCore_go___spec__6___closed__2; @@ -375,6 +373,7 @@ lean_object* l_Lean_MessageData_ofExpr(lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_tryTheoremCore_go___lambda__4___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_tryTheoremWithExtraArgs_x3f___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_Meta_Linear_isLinearTerm(lean_object*); +lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNode___spec__1(size_t, size_t, lean_object*); static lean_object* l_Lean_Meta_Simp_postSEval___lambda__1___closed__1; uint8_t l_Lean_LocalDecl_isImplementationDetail(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_drewritePost(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -388,6 +387,7 @@ LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_synthesizeArgs___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*); static lean_object* l_Lean_Meta_ppOrigin___at_Lean_Meta_Simp_discharge_x3f_x27___spec__1___closed__6; static lean_object* l_Lean_Meta_ppOrigin___at_Lean_Meta_Simp_discharge_x3f_x27___spec__1___closed__1; +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_Simp_discharge_x3f_x27___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeRevMAux___at_Lean_Meta_Simp_dischargeUsingAssumption_x3f___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_Meta_Linear_Nat_simpExpr_x3f(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*); @@ -458,7 +458,6 @@ static lean_object* l___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_tryTheoremCore_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_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_sevalGround___spec__1___closed__2; static lean_object* l___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_tryTheoremCore_go___lambda__3___closed__4; -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_Simp_discharge_x3f_x27___spec__11(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_Array_findSomeRevM_x3f_find___at_Lean_Meta_Simp_dischargeUsingAssumption_x3f___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_simpUsingDecide___lambda__1___closed__14; @@ -481,7 +480,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Simp_DischargeResult_toCtorIdx(uint8_t); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_drewritePre___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_Meta_Simp_sevalGround___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*); static lean_object* l_Lean_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___lambda__3___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_Simp_discharge_x3f_x27___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___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_tryTheoremCore_go___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*); static lean_object* l_Lean_Meta_Simp_mkSEvalContext___closed__2; static lean_object* l_Lean_Meta_Simp_simpUsingDecide___lambda__1___closed__4; @@ -537,6 +535,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Simp_mkSEvalMethods___rarg___boxed(lean_obj lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simpArith(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_preSEval(lean_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_Simp_discharge_x3f_x27___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_dischargeDefault_x3f___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); uint32_t lean_uint32_add(uint32_t, uint32_t); @@ -587,7 +586,6 @@ LEAN_EXPORT lean_object* l_Lean_LocalContext_findDeclRevM_x3f___at_Lean_Meta_Sim static lean_object* l_Lean_Meta_ppSimpTheorem___at___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_tryTheoremCore_go___spec__1___closed__3; lean_object* l_Lean_Expr_mvarId_x21(lean_object*); static lean_object* l_Lean_Meta_Simp_rewrite_x3f___closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_Simp_discharge_x3f_x27___spec__10(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_mkConstWithLevelParams___at_Lean_Meta_Simp_discharge_x3f_x27___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_mkSEvalContext___closed__1; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Meta_Simp_simpMatchDiscrs_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*); @@ -1635,125 +1633,6 @@ return x_41; } } } -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_Simp_discharge_x3f_x27___spec__10(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; uint8_t x_17; -x_14 = lean_st_ref_take(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 = !lean_is_exclusive(x_15); -if (x_17 == 0) -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; size_t x_21; size_t x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; -x_18 = lean_ctor_get(x_15, 3); -x_19 = l_Lean_PersistentArray_toArray___rarg(x_18); -x_20 = lean_array_get_size(x_19); -x_21 = lean_usize_of_nat(x_20); -lean_dec(x_20); -x_22 = 0; -x_23 = l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___spec__1(x_21, x_22, x_19); -x_24 = lean_alloc_ctor(9, 3, 1); -lean_ctor_set(x_24, 0, x_2); -lean_ctor_set(x_24, 1, x_4); -lean_ctor_set(x_24, 2, x_23); -lean_ctor_set_uint8(x_24, sizeof(void*)*3, x_5); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_3); -lean_ctor_set(x_25, 1, x_24); -x_26 = l_Lean_PersistentArray_push___rarg(x_1, x_25); -lean_ctor_set(x_15, 3, x_26); -x_27 = lean_st_ref_set(x_12, x_15, x_16); -x_28 = !lean_is_exclusive(x_27); -if (x_28 == 0) -{ -lean_object* x_29; lean_object* x_30; -x_29 = lean_ctor_get(x_27, 0); -lean_dec(x_29); -x_30 = lean_box(0); -lean_ctor_set(x_27, 0, x_30); -return x_27; -} -else -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_27, 1); -lean_inc(x_31); -lean_dec(x_27); -x_32 = lean_box(0); -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 -{ -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; size_t x_43; size_t x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_34 = lean_ctor_get(x_15, 0); -x_35 = lean_ctor_get(x_15, 1); -x_36 = lean_ctor_get(x_15, 2); -x_37 = lean_ctor_get(x_15, 3); -x_38 = lean_ctor_get(x_15, 4); -x_39 = lean_ctor_get(x_15, 5); -x_40 = lean_ctor_get(x_15, 6); -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_inc(x_34); -lean_dec(x_15); -x_41 = l_Lean_PersistentArray_toArray___rarg(x_37); -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 = l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___spec__1(x_43, x_44, x_41); -x_46 = lean_alloc_ctor(9, 3, 1); -lean_ctor_set(x_46, 0, x_2); -lean_ctor_set(x_46, 1, x_4); -lean_ctor_set(x_46, 2, x_45); -lean_ctor_set_uint8(x_46, sizeof(void*)*3, x_5); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_3); -lean_ctor_set(x_47, 1, x_46); -x_48 = l_Lean_PersistentArray_push___rarg(x_1, x_47); -x_49 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_49, 0, x_34); -lean_ctor_set(x_49, 1, x_35); -lean_ctor_set(x_49, 2, x_36); -lean_ctor_set(x_49, 3, x_48); -lean_ctor_set(x_49, 4, x_38); -lean_ctor_set(x_49, 5, x_39); -lean_ctor_set(x_49, 6, x_40); -x_50 = lean_st_ref_set(x_12, x_49, x_16); -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 = lean_box(0); -if (lean_is_scalar(x_52)) { - x_54 = lean_alloc_ctor(0, 2, 0); -} else { - x_54 = x_52; -} -lean_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_54, 1, x_51); -return x_54; -} -} -} LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___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, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { @@ -1761,76 +1640,269 @@ uint8_t x_14; x_14 = !lean_is_exclusive(x_11); if (x_14 == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; size_t x_23; size_t x_24; lean_object* x_25; lean_object* x_26; 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_15 = lean_ctor_get(x_11, 5); x_16 = l_Lean_replaceRef(x_3, x_15); lean_dec(x_15); lean_ctor_set(x_11, 5, x_16); -x_17 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_4, x_9, x_10, x_11, x_12, x_13); +x_17 = lean_st_ref_get(x_12, x_13); x_18 = lean_ctor_get(x_17, 0); lean_inc(x_18); x_19 = lean_ctor_get(x_17, 1); lean_inc(x_19); lean_dec(x_17); -x_20 = l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_Simp_discharge_x3f_x27___spec__10(x_1, x_2, x_3, x_18, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_19); +x_20 = lean_ctor_get(x_18, 3); +lean_inc(x_20); +lean_dec(x_18); +x_21 = l_Lean_PersistentArray_toArray___rarg(x_20); +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___private_Lean_Util_Trace_0__Lean_addTraceNode___spec__1(x_23, x_24, x_21); +x_26 = lean_alloc_ctor(9, 3, 1); +lean_ctor_set(x_26, 0, x_2); +lean_ctor_set(x_26, 1, x_4); +lean_ctor_set(x_26, 2, x_25); +lean_ctor_set_uint8(x_26, sizeof(void*)*3, x_5); +x_27 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_26, x_9, x_10, x_11, x_12, x_19); lean_dec(x_11); -return x_20; +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_27, 1); +lean_inc(x_29); +lean_dec(x_27); +x_30 = lean_st_ref_take(x_12, x_29); +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 = !lean_is_exclusive(x_31); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; +x_34 = lean_ctor_get(x_31, 3); +lean_dec(x_34); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_3); +lean_ctor_set(x_35, 1, x_28); +x_36 = l_Lean_PersistentArray_push___rarg(x_1, x_35); +lean_ctor_set(x_31, 3, x_36); +x_37 = lean_st_ref_set(x_12, x_31, x_32); +x_38 = !lean_is_exclusive(x_37); +if (x_38 == 0) +{ +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_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_21 = lean_ctor_get(x_11, 0); -x_22 = lean_ctor_get(x_11, 1); -x_23 = lean_ctor_get(x_11, 2); -x_24 = lean_ctor_get(x_11, 3); -x_25 = lean_ctor_get(x_11, 4); -x_26 = lean_ctor_get(x_11, 5); -x_27 = lean_ctor_get(x_11, 6); -x_28 = lean_ctor_get(x_11, 7); -x_29 = lean_ctor_get(x_11, 8); -x_30 = lean_ctor_get(x_11, 9); -x_31 = lean_ctor_get(x_11, 10); -x_32 = lean_ctor_get_uint8(x_11, sizeof(void*)*11); -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_inc(x_21); +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 +{ +lean_object* x_44; lean_object* x_45; 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_44 = lean_ctor_get(x_31, 0); +x_45 = lean_ctor_get(x_31, 1); +x_46 = lean_ctor_get(x_31, 2); +x_47 = lean_ctor_get(x_31, 4); +x_48 = lean_ctor_get(x_31, 5); +x_49 = lean_ctor_get(x_31, 6); +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_31); +x_50 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_50, 0, x_3); +lean_ctor_set(x_50, 1, x_28); +x_51 = l_Lean_PersistentArray_push___rarg(x_1, x_50); +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_47); +lean_ctor_set(x_52, 5, x_48); +lean_ctor_set(x_52, 6, x_49); +x_53 = lean_st_ref_set(x_12, x_52, x_32); +x_54 = lean_ctor_get(x_53, 1); +lean_inc(x_54); +if (lean_is_exclusive(x_53)) { + lean_ctor_release(x_53, 0); + lean_ctor_release(x_53, 1); + x_55 = x_53; +} else { + lean_dec_ref(x_53); + x_55 = lean_box(0); +} +x_56 = lean_box(0); +if (lean_is_scalar(x_55)) { + x_57 = lean_alloc_ctor(0, 2, 0); +} else { + x_57 = x_55; +} +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_54); +return x_57; +} +} +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; uint8_t x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; size_t x_78; size_t x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; +x_58 = lean_ctor_get(x_11, 0); +x_59 = lean_ctor_get(x_11, 1); +x_60 = lean_ctor_get(x_11, 2); +x_61 = lean_ctor_get(x_11, 3); +x_62 = lean_ctor_get(x_11, 4); +x_63 = lean_ctor_get(x_11, 5); +x_64 = lean_ctor_get(x_11, 6); +x_65 = lean_ctor_get(x_11, 7); +x_66 = lean_ctor_get(x_11, 8); +x_67 = lean_ctor_get(x_11, 9); +x_68 = lean_ctor_get(x_11, 10); +x_69 = lean_ctor_get_uint8(x_11, sizeof(void*)*11); +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_dec(x_11); -x_33 = l_Lean_replaceRef(x_3, x_26); -lean_dec(x_26); -x_34 = lean_alloc_ctor(0, 11, 1); -lean_ctor_set(x_34, 0, x_21); -lean_ctor_set(x_34, 1, x_22); -lean_ctor_set(x_34, 2, x_23); -lean_ctor_set(x_34, 3, x_24); -lean_ctor_set(x_34, 4, x_25); -lean_ctor_set(x_34, 5, x_33); -lean_ctor_set(x_34, 6, x_27); -lean_ctor_set(x_34, 7, x_28); -lean_ctor_set(x_34, 8, x_29); -lean_ctor_set(x_34, 9, x_30); -lean_ctor_set(x_34, 10, x_31); -lean_ctor_set_uint8(x_34, sizeof(void*)*11, x_32); -x_35 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_4, x_9, x_10, x_34, x_12, x_13); -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_Meta_Simp_discharge_x3f_x27___spec__10(x_1, x_2, x_3, x_36, x_5, x_6, x_7, x_8, x_9, x_10, x_34, x_12, x_37); -lean_dec(x_34); -return x_38; +x_70 = l_Lean_replaceRef(x_3, x_63); +lean_dec(x_63); +x_71 = lean_alloc_ctor(0, 11, 1); +lean_ctor_set(x_71, 0, x_58); +lean_ctor_set(x_71, 1, x_59); +lean_ctor_set(x_71, 2, x_60); +lean_ctor_set(x_71, 3, x_61); +lean_ctor_set(x_71, 4, x_62); +lean_ctor_set(x_71, 5, x_70); +lean_ctor_set(x_71, 6, x_64); +lean_ctor_set(x_71, 7, x_65); +lean_ctor_set(x_71, 8, x_66); +lean_ctor_set(x_71, 9, x_67); +lean_ctor_set(x_71, 10, x_68); +lean_ctor_set_uint8(x_71, sizeof(void*)*11, x_69); +x_72 = lean_st_ref_get(x_12, x_13); +x_73 = lean_ctor_get(x_72, 0); +lean_inc(x_73); +x_74 = lean_ctor_get(x_72, 1); +lean_inc(x_74); +lean_dec(x_72); +x_75 = lean_ctor_get(x_73, 3); +lean_inc(x_75); +lean_dec(x_73); +x_76 = l_Lean_PersistentArray_toArray___rarg(x_75); +x_77 = lean_array_get_size(x_76); +x_78 = lean_usize_of_nat(x_77); +lean_dec(x_77); +x_79 = 0; +x_80 = l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNode___spec__1(x_78, x_79, x_76); +x_81 = lean_alloc_ctor(9, 3, 1); +lean_ctor_set(x_81, 0, x_2); +lean_ctor_set(x_81, 1, x_4); +lean_ctor_set(x_81, 2, x_80); +lean_ctor_set_uint8(x_81, sizeof(void*)*3, x_5); +x_82 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_81, x_9, x_10, x_71, x_12, x_74); +lean_dec(x_71); +x_83 = lean_ctor_get(x_82, 0); +lean_inc(x_83); +x_84 = lean_ctor_get(x_82, 1); +lean_inc(x_84); +lean_dec(x_82); +x_85 = lean_st_ref_take(x_12, x_84); +x_86 = lean_ctor_get(x_85, 0); +lean_inc(x_86); +x_87 = lean_ctor_get(x_85, 1); +lean_inc(x_87); +lean_dec(x_85); +x_88 = lean_ctor_get(x_86, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_86, 1); +lean_inc(x_89); +x_90 = lean_ctor_get(x_86, 2); +lean_inc(x_90); +x_91 = lean_ctor_get(x_86, 4); +lean_inc(x_91); +x_92 = lean_ctor_get(x_86, 5); +lean_inc(x_92); +x_93 = lean_ctor_get(x_86, 6); +lean_inc(x_93); +if (lean_is_exclusive(x_86)) { + lean_ctor_release(x_86, 0); + lean_ctor_release(x_86, 1); + lean_ctor_release(x_86, 2); + lean_ctor_release(x_86, 3); + lean_ctor_release(x_86, 4); + lean_ctor_release(x_86, 5); + lean_ctor_release(x_86, 6); + x_94 = x_86; +} else { + lean_dec_ref(x_86); + x_94 = lean_box(0); +} +x_95 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_95, 0, x_3); +lean_ctor_set(x_95, 1, x_83); +x_96 = l_Lean_PersistentArray_push___rarg(x_1, x_95); +if (lean_is_scalar(x_94)) { + x_97 = lean_alloc_ctor(0, 7, 0); +} else { + x_97 = x_94; +} +lean_ctor_set(x_97, 0, x_88); +lean_ctor_set(x_97, 1, x_89); +lean_ctor_set(x_97, 2, x_90); +lean_ctor_set(x_97, 3, x_96); +lean_ctor_set(x_97, 4, x_91); +lean_ctor_set(x_97, 5, x_92); +lean_ctor_set(x_97, 6, x_93); +x_98 = lean_st_ref_set(x_12, x_97, x_87); +x_99 = lean_ctor_get(x_98, 1); +lean_inc(x_99); +if (lean_is_exclusive(x_98)) { + lean_ctor_release(x_98, 0); + lean_ctor_release(x_98, 1); + x_100 = x_98; +} else { + lean_dec_ref(x_98); + x_100 = lean_box(0); +} +x_101 = lean_box(0); +if (lean_is_scalar(x_100)) { + x_102 = lean_alloc_ctor(0, 2, 0); +} else { + x_102 = x_100; } +lean_ctor_set(x_102, 0, x_101); +lean_ctor_set(x_102, 1, x_99); +return x_102; } } -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_Simp_discharge_x3f_x27___spec__11(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +} +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_Simp_discharge_x3f_x27___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: { if (lean_obj_tag(x_1) == 0) @@ -1931,7 +2003,7 @@ x_16 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_Simp_disc x_17 = lean_ctor_get(x_16, 1); lean_inc(x_17); lean_dec(x_16); -x_18 = l_MonadExcept_ofExcept___at_Lean_Meta_Simp_discharge_x3f_x27___spec__11(x_5, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_17); +x_18 = l_MonadExcept_ofExcept___at_Lean_Meta_Simp_discharge_x3f_x27___spec__10(x_5, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_17); lean_dec(x_13); lean_dec(x_9); return x_18; @@ -2210,7 +2282,7 @@ 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_Meta_Simp_discharge_x3f_x27___spec__11(x_24, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_34); +x_35 = l_MonadExcept_ofExcept___at_Lean_Meta_Simp_discharge_x3f_x27___spec__10(x_24, 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); @@ -2295,7 +2367,7 @@ 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_Meta_Simp_discharge_x3f_x27___spec__11(x_24, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_54); +x_55 = l_MonadExcept_ofExcept___at_Lean_Meta_Simp_discharge_x3f_x27___spec__10(x_24, 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); @@ -3955,23 +4027,6 @@ lean_dec(x_1); return x_7; } } -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_Simp_discharge_x3f_x27___spec__10___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: -{ -uint8_t x_14; lean_object* x_15; -x_14 = lean_unbox(x_5); -lean_dec(x_5); -x_15 = l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_Simp_discharge_x3f_x27___spec__10(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); -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_15; -} -} LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___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: { @@ -3988,11 +4043,11 @@ lean_dec(x_6); return x_15; } } -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_Simp_discharge_x3f_x27___spec__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_Simp_discharge_x3f_x27___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) { _start: { lean_object* x_10; -x_10 = l_MonadExcept_ofExcept___at_Lean_Meta_Simp_discharge_x3f_x27___spec__11(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_10 = l_MonadExcept_ofExcept___at_Lean_Meta_Simp_discharge_x3f_x27___spec__10(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); diff --git a/stage0/stdlib/Lean/Parser/Term.c b/stage0/stdlib/Lean/Parser/Term.c index 92b9be0dbce3..e7e8d0ce20cc 100644 --- a/stage0/stdlib/Lean/Parser/Term.c +++ b/stage0/stdlib/Lean/Parser/Term.c @@ -31,7 +31,6 @@ lean_object* l_Lean_Parser_checkPrec(lean_object*); static lean_object* l_Lean_Parser_Term_letIdLhs___closed__9; LEAN_EXPORT lean_object* l_Lean_Parser_Term_binderTactic_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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_Termination_terminationBy___closed__20; static lean_object* l_Lean_Parser_Tactic_tacticSeqIndentGt_formatter___closed__5; static lean_object* l_Lean_Parser_Term_withAnonymousAntiquot___closed__3; static lean_object* l_Lean_Parser_Term_let__fun___closed__10; @@ -87,10 +86,10 @@ static lean_object* l_Lean_Parser_Tactic_tacticSeqIndentGt___closed__4; static lean_object* l_Lean_Parser_Term_optExprPrecedence_formatter___closed__3; static lean_object* l_Lean_Parser_Term_withDeclName___closed__7; static lean_object* l_Lean_Parser_Term_letDecl_parenthesizer___closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__16; static lean_object* l_Lean_Parser_Term_matchExpr_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Term_prop_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_structInstLVal___closed__7; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__45; static lean_object* l_Lean_Parser_Term_char_formatter___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_whereDecls_formatter(lean_object*); static lean_object* l_Lean_Parser_Tactic_sepByIndentSemicolon_formatter___closed__2; @@ -169,6 +168,7 @@ static lean_object* l_Lean_Parser_Term_ellipsis___closed__6; static lean_object* l_Lean_Parser_Term_letRecDecl___closed__11; static lean_object* l_Lean_Parser_Term_typeOf___closed__9; static lean_object* l_Lean_Parser_Term_structInstField___closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__12; lean_object* l_Lean_Parser_registerBuiltinParserAttribute(lean_object*, lean_object*, uint8_t, lean_object*); static lean_object* l_Lean_Parser_Term_pipeCompletion___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_byTactic_docString(lean_object*); @@ -177,10 +177,10 @@ static lean_object* l_Lean_Parser_Term_quotedName___closed__2; static lean_object* l_Lean_Parser_Term_rightact_formatter___closed__3; lean_object* l_Lean_Syntax_MonadTraverser_getCur___at_Lean_PrettyPrinter_Parenthesizer_visitArgs___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_unsafe_formatter___closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__54; static lean_object* l_Lean_Parser_Term_match_formatter___closed__4; static lean_object* l_Lean_Parser_Term_structInst_formatter___closed__19; static lean_object* l_Lean_Parser_semicolonOrLinebreak_formatter___closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__31; 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; @@ -196,6 +196,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_binderTactic_parenthesizer__ static lean_object* l_Lean_Parser_Term_unop___closed__6; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_forInMacro_declRange(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_dynamicQuot_declRange___closed__6; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__30; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_typeOf_declRange(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_binrel_formatter___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_namedPattern(lean_object*); @@ -255,10 +256,10 @@ static lean_object* l_Lean_Parser_Term_fun___closed__8; static lean_object* l_Lean_Parser_Term_inaccessible___closed__2; static lean_object* l_Lean_Parser_Term_dotIdent___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_let__fun_docString___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__42; static lean_object* l_Lean_Parser_Term_letExpr___closed__16; static lean_object* l_Lean_Parser_Term_haveI___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Term_attrInstance_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__51; static lean_object* l_Lean_Parser_Term_letDecl___closed__3; lean_object* l_Lean_addBuiltinDocString(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_unsafe_formatter___closed__4; @@ -287,7 +288,6 @@ static lean_object* l_Lean_Parser_Term_bracketedBinder___closed__3; static lean_object* l_Lean_Parser_Term_structInstFieldAbbrev___closed__7; static lean_object* l_Lean_Parser_Term_fun___closed__2; static lean_object* l_Lean_Parser_Tactic_sepByIndentSemicolon___closed__13; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__30; static lean_object* l___regBuiltin_Lean_Parser_Term_noindex_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_ellipsis___closed__8; static lean_object* l_Lean_Parser_Term_letMVar___closed__2; @@ -297,7 +297,6 @@ static lean_object* l_Lean_Parser_Term_macroLastArg_formatter___closed__1; static lean_object* l_Lean_Parser_Term_trailing__parser_formatter___closed__4; static lean_object* l_Lean_Parser_Term_binrel__no__prop___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_unreachable_declRange___closed__6; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__10; static lean_object* l___regBuiltin_Lean_Parser_Term_binop__lazy_formatter___closed__2; static lean_object* l_Lean_Parser_Term_anonymousCtor___closed__5; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_66____closed__1; @@ -348,7 +347,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_nomatch_declRange___closed__ LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_tacticSeqBracketed_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Termination_suffix___closed__8; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_binop__lazy_formatter(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__53; static lean_object* l_Lean_Parser_Command_docComment___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_let__tmp_formatter___closed__2; static lean_object* l_Lean_Parser_Term_clear___closed__2; @@ -430,10 +428,12 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_scoped_parenthesizer(le static lean_object* l___regBuiltin_Lean_Parser_Term_ellipsis_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_scoped_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_macroDollarArg___closed__6; +static lean_object* l_Lean_Parser_Term_hole___closed__8; static lean_object* l_Lean_Parser_Term_dynamicQuot___closed__9; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Tactic_quotSeq(lean_object*); static lean_object* l_Lean_Parser_Term_letIdDecl_formatter___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_typeOf_declRange___closed__3; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__50; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_showTermElabImpl_declRange(lean_object*); static lean_object* l_Lean_Parser_Term_explicit___closed__10; static lean_object* l___regBuiltin_Lean_Parser_Term_forInMacro_x27_declRange___closed__1; @@ -517,7 +517,6 @@ static lean_object* l_Lean_Parser_Term_withAnonymousAntiquot___closed__10; static lean_object* l_Lean_Parser_Term_structInst___closed__7; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Tactic_tacticSeqIndentGt_formatter(lean_object*); static lean_object* l_Lean_Parser_Term_letDecl___closed__10; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__35; static lean_object* l_Lean_Parser_Term_haveDecl_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_waitIfTypeContainsMVar_formatter___closed__2; static lean_object* l_Lean_Parser_Term_matchAltsWhereDecls_formatter___closed__2; @@ -546,7 +545,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_instCoeTSyntaxConsSyntaxNodeKindMkSt static lean_object* l_Lean_Parser_Term_structInstField_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Tactic_tacticSeq1Indented___closed__4; static lean_object* l_Lean_Parser_Command_docComment_formatter___closed__8; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__39; LEAN_EXPORT lean_object* l_Lean_Parser_Term_app_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Term_paren_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Term_optExprPrecedence_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -575,8 +573,8 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_byTactic_x27_parenthesizer__ 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; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__32; 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_5127____closed__40; 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*); @@ -601,12 +599,10 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_quotedName_parenthesizer___c static lean_object* l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Term___hyg_154____closed__13; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_binrel(lean_object*); static lean_object* l_Lean_Parser_Term_structInstArrayRef_parenthesizer___closed__3; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__20; static lean_object* l_Lean_Parser_Term_explicitUniv___closed__11; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_matchDiscr_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Termination_decreasingBy___closed__1; static lean_object* l_Lean_Parser_Term_binderTactic___closed__5; -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127_(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_paren_formatter___closed__2; static lean_object* l_Lean_Parser_Term_structInstField_formatter___closed__6; static lean_object* l_Lean_Parser_Term_pipeCompletion___closed__1; @@ -695,6 +691,8 @@ 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_initFn____x40_Lean_Parser_Term___hyg_5126____closed__42; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__13; static lean_object* l_Lean_Parser_Term_omission___closed__5; static lean_object* l_Lean_Parser_Term_let__fun___closed__14; static lean_object* l_Lean_Parser_Term_dotIdent___closed__2; @@ -718,7 +716,6 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_binrel__no__prop_parent static lean_object* l_Lean_Parser_Term_structInstField_formatter___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_have_formatter___closed__1; static lean_object* l_Lean_Parser_Term_strictImplicitLeftBracket_formatter___closed__3; -static lean_object* l_Lean_Parser_Termination_terminationBy_parenthesizer___closed__11; static lean_object* l_Lean_Parser_Term_assert_formatter___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_scoped_formatter___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Term_binderDefault; @@ -741,6 +738,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_explicitBinder_parenthesizer___boxed static lean_object* l_Lean_Parser_Termination_terminationBy_formatter___closed__1; static lean_object* l_Lean_Parser_Term_showTermElabImpl___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Term_matchExpr_declRange___closed__7; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__11; 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; @@ -863,6 +861,7 @@ static lean_object* l_Lean_Parser_Term_tuple___closed__12; static lean_object* l_Lean_Parser_Term_attrKind___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Term_panic_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_match_formatter___closed__13; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_let__fun_docString(lean_object*); static lean_object* l_Lean_Parser_Term_haveIdDecl_formatter___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Termination_terminationBy_x3f_parenthesizer___closed__1; @@ -1007,10 +1006,8 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_letExpr_formatter(lean_object*, lean static lean_object* l_Lean_Parser_Term_showTermElabImpl_formatter___closed__4; static lean_object* l_Lean_Parser_Term_showTermElabImpl___closed__3; static lean_object* l_Lean_Parser_Term_dynamicQuot_parenthesizer___closed__7; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__23; static lean_object* l_Lean_Parser_Term_generalizingParam_parenthesizer___closed__8; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_typeSpec_parenthesizer(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__26; static lean_object* l_Lean_Parser_Term_waitIfTypeMVar_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_have_formatter___closed__2; lean_object* l_Lean_Parser_checkLinebreakBefore(lean_object*); @@ -1135,6 +1132,7 @@ static lean_object* l_Lean_Parser_Term_letMVar___closed__12; static lean_object* l___regBuiltin_Lean_Parser_Term_showTermElabImpl_declRange___closed__1; static lean_object* l_Lean_Parser_Term_structInstLVal_formatter___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_explicit_declRange___closed__4; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__37; static lean_object* l_Lean_Parser_Term_assert___closed__10; static lean_object* l___regBuiltin_Lean_Parser_Term_type_declRange___closed__4; static lean_object* l_Lean_Parser_Term_binderDefault___closed__2; @@ -1153,6 +1151,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_scoped_parenthesizer___close LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_scoped_formatter(lean_object*); static lean_object* l_Lean_Parser_Term_explicit_parenthesizer___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_optEllipsis_formatter(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__52; static lean_object* l___regBuiltin_Lean_Parser_Term_sufficesDecl_formatter___closed__1; static lean_object* l_Lean_Parser_Term_matchExpr___closed__6; static lean_object* l_Lean_Parser_Term_structInstField___closed__3; @@ -1186,7 +1185,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_app_declRange___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Term_funBinder_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_forall___closed__6; static lean_object* l_Lean_Parser_Term_let__fun_formatter___closed__7; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__37; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_str_declRange(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_letMVar_declRange___closed__6; static lean_object* l_Lean_Parser_Term_noImplicitLambda___closed__9; @@ -1283,7 +1281,6 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_paren_parenthesizer(lea static lean_object* l_Lean_Parser_Term_attrKind___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_letI_docString(lean_object*); static lean_object* l_Lean_Parser_Term_ensureExpectedType___closed__8; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__29; static lean_object* l_Lean_Parser_Term_binrel__no__prop___closed__9; LEAN_EXPORT lean_object* l_Lean_Parser_Term_explicit_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_letExpr_formatter___closed__1; @@ -1329,7 +1326,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_num___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_doubleQuotedName_docString(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Term_sorry; static lean_object* l_Lean_Parser_Term_unsafe_formatter___closed__3; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_ensureTypeOf_formatter___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_ident_declRange___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Term_typeOf; @@ -1366,7 +1362,6 @@ 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_Tactic_tacticSeq1Indented_parenthesizer___closed__3; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__44; static lean_object* l_Lean_Parser_Term_inaccessible_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Term_namedArgument___closed__9; static lean_object* l_Lean_Parser_Term_noImplicitLambda_formatter___closed__1; @@ -1375,6 +1370,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___regBuiltin_Lean_Parser_Term_haveI_formatter___closed__1; static lean_object* l_Lean_Parser_Term_attrKind_formatter___closed__4; @@ -1417,6 +1413,7 @@ static lean_object* l_Lean_Parser_Term_showTermElabImpl___closed__4; static lean_object* l_Lean_Parser_Term_hole___closed__2; static lean_object* l_Lean_Parser_Term_match_parenthesizer___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_let_docString(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__43; static lean_object* l_Lean_Parser_Term_optEllipsis_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_funBinder___closed__1; static lean_object* l_Lean_Parser_Term_waitIfTypeMVar_formatter___closed__5; @@ -1573,6 +1570,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Tactic_tacticSeqBracketed_pa LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_hole_docString(lean_object*); static lean_object* l_Lean_Parser_Tactic_tacticSeqIndentGt_formatter___closed__4; static lean_object* l_Lean_Parser_Term_paren___closed__4; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__22; 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; @@ -1589,7 +1587,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_noErrorIfUnused_declRange___ LEAN_EXPORT lean_object* l_Lean_Parser_Term_scientific; static lean_object* l_Lean_Parser_Term_binop_parenthesizer___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_unsafe_parenthesizer___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__25; lean_object* l_Lean_Parser_sepByIndent_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_waitIfContainsMVar_parenthesizer___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_syntheticHole_declRange(lean_object*); @@ -1794,7 +1791,6 @@ 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_initFn____x40_Lean_Parser_Term___hyg_5127____closed__11; static lean_object* l_Lean_Parser_Term_haveDecl_parenthesizer___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Term_nofun_declRange___closed__4; static lean_object* l_Lean_Parser_Term_tuple___closed__15; @@ -1833,7 +1829,6 @@ static lean_object* l_Lean_Parser_Command_docComment___closed__1; static lean_object* l_Lean_Parser_Term_binop__lazy___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_binop_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_nofun_declRange___closed__7; -static lean_object* l_Lean_Parser_Termination_terminationBy___closed__17; static lean_object* l_Lean_Parser_Term_ellipsis___closed__9; static lean_object* l___regBuiltin_Lean_Parser_Term_tuple_declRange___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Term_leading__parser_formatter___closed__2; @@ -1853,7 +1848,6 @@ static lean_object* l_Lean_Parser_Term_generalizingParam_parenthesizer___closed_ lean_object* l_Lean_Parser_withPosition(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Term_binderIdent_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_forall_declRange___closed__2; -static lean_object* l_Lean_Parser_Termination_terminationBy_formatter___closed__11; static lean_object* l_Lean_Parser_Term_optExprPrecedence___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_show_formatter(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_char___closed__1; @@ -1885,7 +1879,6 @@ static lean_object* l_Lean_Parser_Tactic_quot_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Term_binrel__no__prop_formatter___closed__3; static lean_object* l_Lean_Parser_Term_type___closed__19; static lean_object* l_Lean_Parser_Tactic_quotSeq___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__3; static lean_object* l_Lean_Parser_Term_pipeProj___closed__2; static lean_object* l_Lean_Parser_Term_explicitUniv___closed__10; static lean_object* l_Lean_Parser_Term_binrel___closed__5; @@ -1898,6 +1891,7 @@ static lean_object* l_Lean_Parser_Term_pipeProj___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Term_rightact_declRange___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Term_structInstField; static lean_object* l_Lean_Parser_Term_ensureTypeOf___closed__4; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_char_declRange___closed__5; static lean_object* l_Lean_Parser_Term_letIdDecl___closed__6; static lean_object* l_Lean_Parser_Term_funImplicitBinder_parenthesizer___closed__6; @@ -1973,7 +1967,7 @@ static lean_object* l_Lean_Parser_Termination_suffix___closed__3; 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___regBuiltin_Lean_Parser_Term_attributes_parenthesizer___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__46; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__15; static lean_object* l_Lean_Parser_Term_macroDollarArg_formatter___closed__3; static lean_object* l_Lean_Parser_Term_funStrictImplicitBinder_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_typeAscription_parenthesizer___closed__2; @@ -2032,6 +2026,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_typeAscription_parenthe static lean_object* l_Lean_Parser_Term_leftact___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Term_optType_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_structInst___closed__12; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__40; static lean_object* l___regBuiltin_Lean_Parser_Term_clear_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_typeOf___closed__2; static lean_object* l_Lean_Parser_Term_motive_formatter___closed__4; @@ -2114,6 +2109,7 @@ static lean_object* l_Lean_Parser_Term_fun___closed__10; static lean_object* l_Lean_Parser_Term_whereDecls_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Term_falseVal_formatter___closed__2; static lean_object* l_Lean_Parser_Term_funStrictImplicitBinder_parenthesizer___closed__6; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__31; static lean_object* l___regBuiltin_Lean_Parser_Term_subst_declRange___closed__3; static lean_object* l_Lean_Parser_Term_show___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Term_haveIdLhs_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -2143,7 +2139,9 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Tactic_quot_declRange(lean_o static lean_object* l_Lean_Parser_Term_waitIfTypeMVar_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Term_letrec_formatter___closed__5; static lean_object* l_Lean_Parser_Term_dotIdent___closed__5; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__7; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_waitIfTypeMVar(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__18; static lean_object* l___regBuiltin_Lean_Parser_Term_borrowed_formatter___closed__1; lean_object* l_ReaderT_bind___at_Lean_PrettyPrinter_Parenthesizer_trailingNode_parenthesizer___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_convParser(lean_object*); @@ -2200,7 +2198,6 @@ static lean_object* l_Lean_Parser_Term_typeAscription_parenthesizer___closed__11 static lean_object* l_Lean_Parser_Term_scoped___closed__4; 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_Term_attrInstance_formatter___closed__3; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__41; static lean_object* l_Lean_Parser_Term_optSemicolon_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_letIdLhs_formatter___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_hole_parenthesizer___closed__2; @@ -2298,7 +2295,6 @@ lean_object* l_Lean_Parser_numLit_parenthesizer(lean_object*, lean_object*, lean static lean_object* l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Term___hyg_154____closed__9; LEAN_EXPORT lean_object* l_Lean_Parser_tacticParser_formatter___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic_tacticSeq_formatter___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__5; static lean_object* l_Lean_Parser_Term_namedArgument___closed__7; static lean_object* l_Lean_Parser_Tactic_tacticSeqBracketed_parenthesizer___closed__8; static lean_object* l_Lean_Parser_Termination_terminationBy_formatter___closed__8; @@ -2307,6 +2303,7 @@ static lean_object* l_Lean_Parser_Term_structInst_formatter___closed__11; static lean_object* l___regBuiltin_Lean_Parser_Term_structInstField_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_panic_declRange___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_matchExprPat_parenthesizer___closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__35; static lean_object* l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Termination_suffix___closed__10; static lean_object* l_Lean_Parser_Term_structInstArrayRef___closed__6; @@ -2347,6 +2344,7 @@ lean_object* l_Lean_PrettyPrinter_Formatter_parserOfStack_formatter___boxed(lean static lean_object* l_Lean_Parser_Term_waitIfTypeContainsMVar___closed__2; static lean_object* l_Lean_Parser_Termination_decreasingBy___closed__8; static lean_object* l___regBuiltin_Lean_Parser_Tactic_seq1_formatter___closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_withDeclName_parenthesizer___closed__2; 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*); @@ -2507,6 +2505,7 @@ 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_initFn____x40_Lean_Parser_Term___hyg_5126____closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Term_nofun; static lean_object* l_Lean_Parser_Term_arrow___closed__5; static lean_object* l_Lean_Parser_Term_let_formatter___closed__5; @@ -2590,6 +2589,7 @@ static lean_object* l_Lean_Parser_Term_letRecDecls___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_unop_declRange___closed__1; static lean_object* l_Lean_Parser_Term_forall___closed__14; static lean_object* l_Lean_Parser_Term_structInstLVal___closed__13; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__29; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_doubleQuotedName_formatter(lean_object*); static lean_object* l_Lean_Parser_Term_show_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Term_omission___closed__3; @@ -2664,6 +2664,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_arrow_parenthesizer___closed static lean_object* l_Lean_Parser_Term_suffices_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Term_whereDecls___closed__9; static lean_object* l___regBuiltin_Lean_Parser_Term_show_declRange___closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__9; static lean_object* l_Lean_Parser_Term_binop___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Term_trailing__parser; static lean_object* l_Lean_Parser_Term_basicFun___closed__4; @@ -2694,7 +2695,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_omission_declRange___closed_ static lean_object* l_Lean_Parser_Term_binderDefault___closed__5; static lean_object* l_Lean_Parser_Term_sorry_formatter___closed__3; static lean_object* l_Lean_Parser_Term_binrel_formatter___closed__4; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__28; static lean_object* l_Lean_Parser_Term_show___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_binop_declRange___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Term_scientific_declRange___closed__6; @@ -2738,7 +2738,6 @@ static lean_object* l_Lean_Parser_Term_letIdBinder_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_fun_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Term_withAnonymousAntiquot___closed__11; static lean_object* l___regBuiltin_Lean_Parser_Term_match_declRange___closed__4; -static lean_object* l_Lean_Parser_Termination_terminationBy_formatter___closed__12; static lean_object* l___regBuiltin_Lean_Parser_Term_dbgTrace_parenthesizer___closed__2; lean_object* l_Lean_Parser_checkNoWsBefore(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Term_char; @@ -2911,6 +2910,7 @@ static lean_object* l_Lean_Parser_Term_let__tmp___closed__2; static lean_object* l_Lean_Parser_Term_matchExprPat_formatter___closed__1; static lean_object* l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Term___hyg_154____closed__20; static lean_object* l_Lean_Parser_Term_namedPattern_formatter___closed__4; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__46; static lean_object* l_Lean_Parser_Term_let__tmp_formatter___closed__6; static lean_object* l_Lean_Parser_Term_structInst_parenthesizer___closed__16; static lean_object* l_Lean_Parser_Term_structInstLVal_parenthesizer___closed__5; @@ -2972,6 +2972,7 @@ lean_object* l_Lean_PrettyPrinter_Formatter_checkColGt_formatter___boxed(lean_ob static lean_object* l_Lean_Parser_Term_motive___closed__11; static lean_object* l___regBuiltin_Lean_Parser_Term_let_declRange___closed__3; static lean_object* l_Lean_Parser_Term_binrel__no__prop_parenthesizer___closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__48; LEAN_EXPORT lean_object* l_Lean_Parser_Term_unop; static lean_object* l_Lean_Parser_Term_showTermElabImpl___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Term_forall_declRange___closed__6; @@ -3027,7 +3028,6 @@ 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; lean_object* l_Lean_PrettyPrinter_Formatter_rawCh_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__14; static lean_object* l_Lean_Parser_Term_sort_formatter___closed__3; lean_object* l_Lean_Parser_many_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_motive_parenthesizer___closed__1; @@ -3050,7 +3050,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_ensureExpectedType_declRange static lean_object* l_Lean_Parser_Termination_terminationBy_formatter___closed__4; static lean_object* l_Lean_Parser_Term_binrel__no__prop_formatter___closed__1; static lean_object* l_Lean_Parser_Tactic_quot___closed__4; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__17; static lean_object* l_Lean_Parser_Term_let__delayed___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_noErrorIfUnused_parenthesizer(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_structInst_declRange___closed__4; @@ -3109,6 +3108,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_leading__parser_formatter___ static lean_object* l_Lean_Parser_Term_matchExpr___closed__7; static lean_object* l_Lean_Parser_Term_ellipsis_formatter___closed__5; static lean_object* l_Lean_Parser_Term_generalizingParam_parenthesizer___closed__10; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__47; static lean_object* l_Lean_Parser_Term_byTactic_x27___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Term_syntheticHole_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_stateRefT_parenthesizer___closed__2; @@ -3121,7 +3121,6 @@ static lean_object* l_Lean_Parser_Term_strictImplicitRightBracket___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_borrowed(lean_object*); static lean_object* l_Lean_Parser_Term_matchExprPat_formatter___closed__2; static lean_object* l_Lean_Parser_Term_strictImplicitLeftBracket___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__34; static lean_object* l___regBuiltin_Lean_Parser_Term_defaultOrOfNonempty_declRange___closed__6; static lean_object* l_Lean_Parser_Term_letMVar___closed__8; static lean_object* l_Lean_Parser_Term_letRecDecl___closed__2; @@ -3143,7 +3142,6 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_dbgTrace_formatter(lean static lean_object* l_Lean_Parser_Term_depArrow___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Term_show; static lean_object* l_Lean_Parser_Term_letrec_formatter___closed__8; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__12; static lean_object* l_Lean_Parser_Term_optIdent___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_matchAltsWhereDecls_parenthesizer___closed__2; uint8_t l_Lean_Syntax_isIdent(lean_object*); @@ -3154,6 +3152,7 @@ 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_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; @@ -3325,6 +3324,7 @@ static lean_object* l_Lean_Parser_Term_unreachable___closed__1; static lean_object* l_Lean_Parser_Term_binrel__no__prop___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_forall_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_waitIfTypeMVar_formatter___closed__6; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__27; static lean_object* l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Term_attrKind___closed__8; static lean_object* l_Lean_Parser_Term_ensureExpectedType___closed__4; @@ -3352,7 +3352,6 @@ static lean_object* l_Lean_Parser_Term_matchAltsWhereDecls_formatter___closed__5 static lean_object* l_Lean_Parser_Term_structInstArrayRef_formatter___closed__5; static lean_object* l_Lean_Parser_Term_dotIdent_parenthesizer___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_showTermElabImpl_docString(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__54; LEAN_EXPORT lean_object* l_Lean_Parser_Term_withDeclName_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_sufficesDecl_formatter___closed__5; static lean_object* l_Lean_Parser_Term_forall_formatter___closed__5; @@ -3410,7 +3409,6 @@ static lean_object* l_Lean_Parser_Term_haveI___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_forInMacro_formatter___closed__1; static lean_object* l_Lean_Parser_Term_ident_formatter___closed__1; lean_object* l_Lean_Parser_ppSpace_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__43; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_paren(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Term_haveEqnsDecl_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_instBinder___closed__4; @@ -3452,6 +3450,7 @@ static lean_object* l_Lean_Parser_Term_motive___closed__3; static lean_object* l_Lean_Parser_Term_waitIfTypeContainsMVar_formatter___closed__1; static lean_object* l_Lean_Parser_Term_letRecDecls___closed__3; static lean_object* l_Lean_Parser_Term_basicFun_formatter___closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__24; static lean_object* l___regBuiltin_Lean_Parser_Term_letI_docString___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_tuple_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Term_let__fun___closed__4; @@ -3574,6 +3573,7 @@ static lean_object* l_Lean_Parser_Term_structInstLVal_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_stateRefT_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_instBinder___closed__5; static lean_object* l_Lean_Parser_Term_byTactic_x27___closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__53; static lean_object* l_Lean_Parser_Term_binderTactic_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Term_letMVar___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Term_binrel__no__prop_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -3624,7 +3624,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_waitIfTypeMVar_declRange___c 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_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__13; static lean_object* l___regBuiltin_Lean_Parser_Term_anonymousCtor_declRange___closed__3; static lean_object* l_Lean_Parser_Term_whereDecls___closed__6; static lean_object* l_Lean_Parser_Term_structInstField___closed__9; @@ -3640,6 +3639,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_macroDollarArg_formatte static lean_object* l_Lean_Parser_Term_unsafe___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_stateRefT_declRange___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_noErrorIfUnused_declRange___closed__3; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__39; static lean_object* l_Lean_Parser_Term_attrKind___closed__6; static lean_object* l_Lean_Parser_Termination_decreasingBy_formatter___closed__6; static lean_object* l_Lean_Parser_Term_generalizingParam_parenthesizer___closed__2; @@ -3796,6 +3796,7 @@ static lean_object* l_Lean_Parser_Term_withDeclName___closed__6; static lean_object* l_Lean_Parser_Term_fromTerm_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Term_stateRefT___closed__10; lean_object* l_Lean_Parser_registerAliasCore___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__10; static lean_object* l___regBuiltin_Lean_Parser_Term_dbgTrace_formatter___closed__1; static lean_object* l_Lean_Parser_Term_matchExprPat_formatter___closed__3; static lean_object* l_Lean_Parser_Term_letMVar_parenthesizer___closed__2; @@ -3850,15 +3851,14 @@ static lean_object* l_Lean_Parser_Term_matchExprPat_formatter___closed__7; static lean_object* l_Lean_Parser_Term_matchAlts___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_forInMacro_x27_declRange___closed__5; static lean_object* l_Lean_Parser_Term_let___closed__3; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__19; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_suffices_declRange(lean_object*); static lean_object* l_Lean_Parser_Term_attributes_parenthesizer___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_instCoeTSyntaxConsSyntaxNodeKindMkStr4Nil__1___boxed(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_namedPattern_parenthesizer(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__48; static lean_object* l_Lean_Parser_Term_assert_formatter___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Term_basicFun_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_forInMacro_declRange___closed__7; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__22; static lean_object* l_Lean_Parser_Term_match_formatter___closed__8; static lean_object* l_Lean_Parser_Term_forall_parenthesizer___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Term_subst_parenthesizer___closed__1; @@ -3914,7 +3914,6 @@ 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; static lean_object* l_Lean_Parser_Term_funImplicitBinder___closed__2; -static lean_object* l_Lean_Parser_Termination_terminationBy___closed__16; static lean_object* l_Lean_Parser_Term_typeAscription___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Term_arrow_formatter___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_noImplicitLambda_declRange(lean_object*); @@ -4084,7 +4083,6 @@ static lean_object* l_Lean_Parser_Term_nomatch_formatter___closed__1; static lean_object* l_Lean_Parser_Term_typeOf___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_structInst(lean_object*); static lean_object* l_Lean_Parser_Term_syntheticHole_parenthesizer___closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__27; lean_object* l_id___rarg___boxed(lean_object*); static lean_object* l_Lean_Parser_Term_subst___closed__8; static lean_object* l_Lean_Parser_Term_byTactic_parenthesizer___closed__6; @@ -4092,6 +4090,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_matchAlt_parenthesizer(lean_object*, static lean_object* l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Term___hyg_154____closed__6; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_app_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Tactic_tacticSeqBracketed_parenthesizer___closed__11; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__3; 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*); @@ -4150,9 +4149,11 @@ static lean_object* l_Lean_Parser_Term_doubleQuotedName___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_withDeclName_docString___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_doubleQuotedName_declRange___closed__4; static lean_object* l_Lean_Parser_Term_match_parenthesizer___closed__4; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__55; 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_Term_structInst___closed__10; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__25; static lean_object* l_Lean_Parser_Term_prop_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Tactic_tacticSeq_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Termination_suffix_formatter___closed__2; @@ -4172,7 +4173,6 @@ static lean_object* l_Lean_Parser_Term_typeAscription___closed__11; static lean_object* l___regBuiltin_Lean_Parser_Term_explicitUniv_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Tactic_seq1_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_structInst_declRange___closed__2; -static lean_object* l_Lean_Parser_Termination_terminationBy___closed__19; static lean_object* l_Lean_Parser_Term_structInst___closed__15; static lean_object* l___regBuiltin_Lean_Parser_Term_paren_declRange___closed__7; static lean_object* l_Lean_Parser_Term_letrec___closed__5; @@ -4229,6 +4229,7 @@ static lean_object* l_Lean_Parser_Term_fun___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Term_namedArgument; static lean_object* l_Lean_Parser_Term_haveI_formatter___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_structInstArrayRef_parenthesizer(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__14; LEAN_EXPORT lean_object* l_Lean_Parser_Term_have_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_dotIdent_declRange___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_hole_declRange___closed__7; @@ -4277,7 +4278,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_paren_parenthesizer(lean_object*, le static lean_object* l_Lean_Parser_Term_attributes_parenthesizer___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_proj_formatter___closed__1; static lean_object* l_Lean_Parser_Termination_terminationBy_x3f_formatter___closed__3; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__47; static lean_object* l___regBuiltin_Lean_Parser_Term_matchExpr_declRange___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Termination_decreasingBy_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_panic_parenthesizer___closed__4; @@ -4307,7 +4307,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_structInstFieldAbbrev_format LEAN_EXPORT lean_object* l_Lean_Parser_Term_optEllipsis_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Term_haveDecl_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_cdot_declRange___closed__5; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__15; LEAN_EXPORT lean_object* l_Lean_Parser_Term_unsafe_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_waitIfTypeContainsMVar___closed__3; static lean_object* l_Lean_Parser_Term_nofun_parenthesizer___closed__3; @@ -4353,7 +4352,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_forall_declRange___closed__7 LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_waitIfTypeMVar_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Term_forInMacro___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_subst_formatter___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__16; LEAN_EXPORT lean_object* l_Lean_Parser_Termination_terminationBy_x3f_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Term_syntheticHole_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_dotIdent_formatter___closed__1; @@ -4382,7 +4380,6 @@ 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_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__18; static lean_object* l___regBuiltin_Lean_Parser_Term_haveI_declRange___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Term_syntheticHole_declRange___closed__7; static lean_object* l_Lean_Parser_Term_cdot_formatter___closed__5; @@ -4405,11 +4402,11 @@ static lean_object* l_Lean_Parser_Term_stateRefT_formatter___closed__1; static lean_object* l_Lean_Parser_Term_falseVal_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_pipeCompletion_declRange___closed__2; static lean_object* l_Lean_Parser_Term_nomatch_parenthesizer___closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__8; static lean_object* l_Lean_Parser_Command_docComment___closed__9; LEAN_EXPORT lean_object* l_Lean_Parser_Term_dotIdent; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_letrec_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Term_structInst_formatter___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__24; 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; @@ -4422,7 +4419,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_macroLastArg; LEAN_EXPORT lean_object* l_Lean_Parser_Term_showTermElabImpl_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_letrec___closed__10; static lean_object* l_Lean_Parser_Tactic_sepByIndentSemicolon___closed__3; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__7; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_withDeclName_formatter(lean_object*); static lean_object* l_Lean_Parser_Term_namedArgument_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_ensureExpectedType_formatter___closed__1; @@ -4466,6 +4462,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_proj_formatter(lean_obj static lean_object* l_Lean_Parser_Tactic_tacticSeqBracketed___closed__2; static lean_object* l_Lean_Parser_Term_funStrictImplicitBinder_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Term_trailing__parser___closed__10; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__21; static lean_object* l_Lean_Parser_Term_paren_parenthesizer___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_unsafe_docString(lean_object*); static lean_object* l_Lean_Parser_Term_clear_parenthesizer___closed__1; @@ -4475,7 +4472,6 @@ static lean_object* l_Lean_Parser_Term_binderTactic_formatter___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_doubleQuotedName_docString___closed__1; static lean_object* l_Lean_Parser_Term_have_formatter___closed__6; static lean_object* l_Lean_Parser_Term_forall___closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__36; LEAN_EXPORT lean_object* l_Lean_Parser_Term_let__tmp; static lean_object* l_Lean_Parser_Tactic_tacticSeq___closed__6; static lean_object* l_Lean_Parser_Term_letRecDecls_parenthesizer___closed__3; @@ -4483,7 +4479,6 @@ static lean_object* l_Lean_Parser_Term_dynamicQuot_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Term_trailing__parser___closed__9; static lean_object* l___regBuiltin_Lean_Parser_Term_structInstFieldAbbrev_formatter___closed__1; 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_initFn____x40_Lean_Parser_Term___hyg_5127____closed__4; static lean_object* l_Lean_Parser_Tactic_quot_parenthesizer___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_explicitUniv_formatter(lean_object*); static lean_object* l_Lean_Parser_Term_structInstFieldAbbrev_formatter___closed__8; @@ -4506,12 +4501,15 @@ static lean_object* l_Lean_Parser_Term_clear_formatter___closed__1; static lean_object* l_Lean_Parser_Tactic_tacticSeqBracketed_formatter___closed__8; static lean_object* l___regBuiltin_Lean_Parser_Term_structInstLVal_formatter___closed__2; static lean_object* l_Lean_Parser_Term_quotedName___closed__4; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__28; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_tacticSeq_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_showTermElabImpl_formatter___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_dynamicQuot_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Term_dbgTrace___closed__13; static lean_object* l_Lean_Parser_Term_namedArgument_parenthesizer___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Term_let__delayed_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__33; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__36; 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_unsafe___closed__8; @@ -4573,6 +4571,7 @@ static lean_object* l_Lean_Parser_Term_ensureTypeOf___closed__3; static lean_object* l_Lean_Parser_Term_namedArgument___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_noErrorIfUnused_formatter___closed__1; lean_object* l_Lean_PrettyPrinter_Parenthesizer_checkStackTop_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__38; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_letrec_declRange(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_syntheticHole(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_unop_parenthesizer___closed__2; @@ -4604,7 +4603,6 @@ static lean_object* l_Lean_Parser_Term_showTermElabImpl_parenthesizer___closed__ static lean_object* l_Lean_Parser_Term_instBinder_formatter___closed__7; static lean_object* l_Lean_Parser_Term_anonymousCtor___closed__4; lean_object* l_Lean_Parser_checkWsBefore(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__32; static lean_object* l_Lean_Parser_Term_let__tmp_formatter___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_let__delayed_declRange___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_str_declRange___closed__1; @@ -4631,7 +4629,6 @@ static lean_object* l_Lean_Parser_Term_hole___closed__4; static lean_object* l_Lean_Parser_Term_haveI_parenthesizer___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_ensureExpectedType_formatter(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Term_ident; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__51; static lean_object* l_Lean_Parser_Term_whereDecls___closed__7; static lean_object* l_Lean_Parser_Term_suffices___closed__6; static lean_object* l_Lean_Parser_Term_haveI___closed__4; @@ -4640,20 +4637,19 @@ static lean_object* l_Lean_Parser_Term_rightact___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_trueVal_parenthesizer___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_generalizingParam_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_attributes___closed__3; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__33; static lean_object* l_Lean_Parser_Term_optEllipsis_formatter___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_generalizingParam_parenthesizer___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_arrow; static lean_object* l_Lean_Parser_Term_letExpr_parenthesizer___closed__9; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_ensureExpectedType(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Term_bracketedBinder_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__21; static lean_object* l_Lean_Parser_Term_withDeclName_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Term_paren_formatter___closed__2; static lean_object* l_Lean_Parser_Term_attributes_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_prop_declRange___closed__1; static lean_object* l_Lean_Parser_Term_letExpr___closed__2; static lean_object* l_Lean_Parser_Term_namedPattern_parenthesizer___closed__6; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__34; static lean_object* l_Lean_Parser_Term_matchExprAlts_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_whereDecls_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Tactic_tacticSeqBracketed_formatter___closed__10; @@ -4796,10 +4792,10 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_panic_parenthesizer(lean_object*, le 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_initFn____x40_Lean_Parser_Term___hyg_5127____closed__19; static lean_object* l_Lean_Parser_Term_structInstArrayRef___closed__11; static lean_object* l_Lean_Parser_Term_binop__lazy_formatter___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_cdot_parenthesizer___closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__20; static lean_object* l_Lean_Parser_Term_structInstFieldAbbrev___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_structInst_formatter___closed__1; static lean_object* l_Lean_Parser_Term_type___closed__11; @@ -4818,6 +4814,7 @@ static lean_object* l_Lean_Parser_Term_anonymousCtor_formatter___closed__3; static lean_object* l_Lean_Parser_Tactic_tacticSeqBracketed_parenthesizer___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Term_let__tmp_declRange___closed__3; static lean_object* l_Lean_Parser_Term_proj_formatter___closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__44; static lean_object* l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_typeAscription___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Term_binderDefault_parenthesizer___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -4835,7 +4832,6 @@ lean_object* l_Lean_Parser_levelParser_parenthesizer(lean_object*, lean_object*, static lean_object* l_Lean_Parser_Term_matchExpr___closed__9; static lean_object* l_Lean_Parser_Term_ellipsis_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_waitIfTypeMVar_parenthesizer___closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_letExpr_declRange___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_app_declRange___closed__1; static lean_object* l_Lean_Parser_Termination_terminationBy_x3f___closed__2; @@ -4858,6 +4854,7 @@ static lean_object* l_Lean_Parser_Term_waitIfTypeMVar_parenthesizer___closed__1; lean_object* l_Lean_PrettyPrinter_Parenthesizer_notFollowedBy_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_defaultOrOfNonempty(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_dotIdent_declRange(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__23; static lean_object* l_Lean_Parser_Term_let_formatter___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_binop__lazy_declRange(lean_object*); static lean_object* l_Lean_Parser_Term_unsafe_formatter___closed__1; @@ -4866,7 +4863,6 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_nomatch_formatter(lean_ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Tactic_seq1_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Term_letIdDecl___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Term_completion_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Termination_terminationBy_parenthesizer___closed__12; static lean_object* l___regBuiltin_Lean_Parser_Term_waitIfTypeContainsMVar_declRange___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_hole_declRange___closed__2; static lean_object* l_Lean_Parser_Term_ensureTypeOf___closed__9; @@ -4941,7 +4937,6 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_borrowed_parenthesizer( static lean_object* l___regBuiltin_Lean_Parser_Term_fun_declRange___closed__1; static lean_object* l_Lean_Parser_semicolonOrLinebreak___closed__1; static lean_object* l_Lean_Parser_Term_generalizingParam___closed__13; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__8; static lean_object* l_Lean_Parser_Term_trueVal___closed__2; static lean_object* l_Lean_Parser_Term_noImplicitLambda_parenthesizer___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Term_dbgTrace; @@ -4992,7 +4987,6 @@ static lean_object* l_Lean_Parser_Term_implicitBinder___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_attrInstance_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_noindex___closed__4; static lean_object* l_Lean_Parser_Term_letPatDecl_parenthesizer___closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__52; 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; @@ -5007,6 +5001,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_prop_formatter(lean_object*, lean_ob static lean_object* l_Lean_Parser_Term_let_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_withDeclName_formatter___closed__3; static lean_object* l_Lean_Parser_Term_matchExpr_formatter___closed__6; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__41; static lean_object* l_Lean_Parser_Term_leading__parser_parenthesizer___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_paren_declRange___closed__5; static lean_object* l_Lean_Parser_Term_matchExpr_formatter___closed__8; @@ -5146,7 +5141,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_arrow_parenthesizer___closed static lean_object* l_Lean_Parser_Term_proj_formatter___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_app_formatter___closed__1; static lean_object* l_Lean_Parser_Term_forInMacro_x27_formatter___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__38; static lean_object* l___regBuiltin_Lean_Parser_Term_ensureExpectedType_formatter___closed__2; static lean_object* l_Lean_Parser_Term_tuple___closed__8; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_unsafe(lean_object*); @@ -5172,7 +5166,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Tactic_tacticSeqBracketed_formatt static lean_object* l___regBuiltin_Lean_Parser_Term_dbgTrace_formatter___closed__2; static lean_object* l_Lean_Parser_Term_typeAscription_parenthesizer___closed__8; static lean_object* l___regBuiltin_Lean_Parser_Term_letMVar_declRange___closed__4; -static lean_object* l_Lean_Parser_Termination_terminationBy___closed__18; static lean_object* l___regBuiltin_Lean_Parser_Term_forInMacro_declRange___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Term_withDeclName_declRange___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_match(lean_object*); @@ -5208,7 +5201,7 @@ static lean_object* l_Lean_Parser_Term_matchExprPat_formatter___closed__6; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_unop_formatter(lean_object*); 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_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__17; static lean_object* l___regBuiltin_Lean_Parser_Term_waitIfTypeMVar_declRange___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_sepByIndentSemicolon(lean_object*); static lean_object* l_Lean_Parser_Term_forall___closed__9; @@ -5229,6 +5222,7 @@ static lean_object* l_Lean_Parser_Term_generalizingParam___closed__1; static lean_object* l_Lean_Parser_Termination_terminationBy___closed__14; static lean_object* l_Lean_Parser_Term_waitIfContainsMVar___closed__9; static lean_object* l_Lean_Parser_Term_omission___closed__1; +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126_(lean_object*); 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; @@ -5313,7 +5307,6 @@ static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_66____clo LEAN_EXPORT lean_object* l_Lean_Parser_Term_noErrorIfUnused_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_typeOf_formatter___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_matchDiscr_parenthesizer___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__55; static lean_object* l_Lean_Parser_Term_forInMacro_formatter___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_stateRefT_declRange___closed__5; static lean_object* l_Lean_Parser_Term_funBinder_formatter___closed__3; @@ -5520,6 +5513,7 @@ extern lean_object* l_Lean_Parser_leadPrec; static lean_object* l_Lean_Parser_Term_ensureTypeOf_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Tactic_tacticSeqIndentGt_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Term_letrec___closed__8; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__26; static lean_object* l___regBuiltin_Lean_Parser_Term_ensureTypeOf_declRange___closed__1; static lean_object* l_Lean_Parser_Term_borrowed___closed__6; static lean_object* l_Lean_Parser_Term_forall_formatter___closed__4; @@ -5527,6 +5521,7 @@ static lean_object* l_Lean_Parser_Term_optExprPrecedence_parenthesizer___closed_ static lean_object* l_Lean_Parser_Term_let__tmp_parenthesizer___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Term_nofun_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_attrKind_parenthesizer___closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__49; static lean_object* l_Lean_Parser_Term_pipeProj___closed__9; static lean_object* l_Lean_Parser_Term_waitIfContainsMVar___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_binderDefault_parenthesizer___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -5601,7 +5596,6 @@ static lean_object* l_Lean_Parser_Term_argument___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Term_let_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_optEllipsis___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_typeAscription_declRange___closed__3; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__49; static lean_object* l_Lean_Parser_Term_depArrow_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_binrel__no__prop_docString___closed__1; static lean_object* l_Lean_Parser_Term_tuple_formatter___closed__7; @@ -5690,12 +5684,13 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_defaultOrOfNonempty_parenthesizer(le static lean_object* l_Lean_Parser_Term_forInMacro_formatter___closed__4; static lean_object* l_Lean_Parser_Term_let___closed__2; static lean_object* l_Lean_Parser_Term_nofun___closed__6; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__2; static lean_object* l_Lean_Parser_Term_letrec_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Term_paren___closed__8; LEAN_EXPORT lean_object* l_Lean_Parser_Term_prop_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_letRecDecls_formatter___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__9; static lean_object* l_Lean_Parser_Term_stateRefT___closed__8; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__45; 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___regBuiltin_Lean_Parser_Term_attrKind_formatter(lean_object*); static lean_object* l_Lean_Parser_Term_type___closed__7; @@ -5704,7 +5699,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_let__fun_declRange___closed_ static lean_object* l_Lean_Parser_Command_docComment___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_binop_declRange___closed__5; static lean_object* l_Lean_Parser_Term_rightact_formatter___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__50; static lean_object* l_Lean_Parser_Term_attributes___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Term_implicitBinder_formatter(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_nofun_declRange___closed__2; @@ -7443,520 +7437,6 @@ x_1 = l_Lean_Parser_semicolonOrLinebreak___closed__3; return x_1; } } -static lean_object* _init_l_Lean_Parser_Termination_terminationBy___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Termination", 11); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Termination_terminationBy___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("terminationBy", 13); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Termination_terminationBy___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_Command_docComment___closed__1; -x_2 = l_Lean_Parser_Command_docComment___closed__2; -x_3 = l_Lean_Parser_Termination_terminationBy___closed__1; -x_4 = l_Lean_Parser_Termination_terminationBy___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_Termination_terminationBy___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; -x_1 = l_Lean_Parser_Termination_terminationBy___closed__2; -x_2 = l_Lean_Parser_Termination_terminationBy___closed__3; -x_3 = 1; -x_4 = 0; -x_5 = l_Lean_Parser_mkAntiquot(x_1, x_2, x_3, x_4); -return x_5; -} -} -static lean_object* _init_l_Lean_Parser_Termination_terminationBy___closed__5() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("termination_by ", 15); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Termination_terminationBy___closed__6() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Termination_terminationBy___closed__5; -x_2 = l_Lean_Parser_symbol(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Parser_Termination_terminationBy___closed__7() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("_", 1); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Termination_terminationBy___closed__8() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Termination_terminationBy___closed__7; -x_2 = l_Lean_Parser_symbol(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Parser_Termination_terminationBy___closed__9() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_ident; -x_2 = l_Lean_Parser_Termination_terminationBy___closed__8; -x_3 = l_Lean_Parser_orelse(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser_Termination_terminationBy___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_Termination_terminationBy___closed__9; -x_3 = l_Lean_Parser_andthen(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser_Termination_terminationBy___closed__11() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Termination_terminationBy___closed__10; -x_2 = l_Lean_Parser_many(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Parser_Termination_terminationBy___closed__12() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Termination_terminationBy___closed__11; -x_2 = l_Lean_Parser_darrow___closed__2; -x_3 = l_Lean_Parser_andthen(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser_Termination_terminationBy___closed__13() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Termination_terminationBy___closed__12; -x_2 = l_Lean_Parser_atomic(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Parser_Termination_terminationBy___closed__14() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Termination_terminationBy___closed__13; -x_2 = l_Lean_Parser_optional(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Parser_Termination_terminationBy___closed__15() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(0u); -x_2 = l_Lean_Parser_termParser(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Parser_Termination_terminationBy___closed__16() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Termination_terminationBy___closed__14; -x_2 = l_Lean_Parser_Termination_terminationBy___closed__15; -x_3 = l_Lean_Parser_andthen(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser_Termination_terminationBy___closed__17() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Termination_terminationBy___closed__6; -x_2 = l_Lean_Parser_Termination_terminationBy___closed__16; -x_3 = l_Lean_Parser_andthen(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser_Termination_terminationBy___closed__18() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Termination_terminationBy___closed__3; -x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Termination_terminationBy___closed__17; -x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); -return x_4; -} -} -static lean_object* _init_l_Lean_Parser_Termination_terminationBy___closed__19() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Termination_terminationBy___closed__4; -x_2 = l_Lean_Parser_Termination_terminationBy___closed__18; -x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser_Termination_terminationBy___closed__20() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Termination_terminationBy___closed__3; -x_2 = l_Lean_Parser_Termination_terminationBy___closed__19; -x_3 = l_Lean_Parser_withCache(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser_Termination_terminationBy() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Parser_Termination_terminationBy___closed__20; -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Termination_terminationBy_x3f___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("terminationBy\?", 14); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Termination_terminationBy_x3f___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_Command_docComment___closed__1; -x_2 = l_Lean_Parser_Command_docComment___closed__2; -x_3 = l_Lean_Parser_Termination_terminationBy___closed__1; -x_4 = l_Lean_Parser_Termination_terminationBy_x3f___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_Termination_terminationBy_x3f___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; -x_1 = l_Lean_Parser_Termination_terminationBy_x3f___closed__1; -x_2 = l_Lean_Parser_Termination_terminationBy_x3f___closed__2; -x_3 = 1; -x_4 = 0; -x_5 = l_Lean_Parser_mkAntiquot(x_1, x_2, x_3, x_4); -return x_5; -} -} -static lean_object* _init_l_Lean_Parser_Termination_terminationBy_x3f___closed__4() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("termination_by\?", 15); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Termination_terminationBy_x3f___closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Termination_terminationBy_x3f___closed__4; -x_2 = l_Lean_Parser_symbol(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Parser_Termination_terminationBy_x3f___closed__6() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Termination_terminationBy_x3f___closed__2; -x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Termination_terminationBy_x3f___closed__5; -x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); -return x_4; -} -} -static lean_object* _init_l_Lean_Parser_Termination_terminationBy_x3f___closed__7() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Termination_terminationBy_x3f___closed__3; -x_2 = l_Lean_Parser_Termination_terminationBy_x3f___closed__6; -x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser_Termination_terminationBy_x3f___closed__8() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Termination_terminationBy_x3f___closed__2; -x_2 = l_Lean_Parser_Termination_terminationBy_x3f___closed__7; -x_3 = l_Lean_Parser_withCache(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser_Termination_terminationBy_x3f() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Parser_Termination_terminationBy_x3f___closed__8; -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Termination_decreasingBy___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("decreasingBy", 12); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Termination_decreasingBy___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_Command_docComment___closed__1; -x_2 = l_Lean_Parser_Command_docComment___closed__2; -x_3 = l_Lean_Parser_Termination_terminationBy___closed__1; -x_4 = l_Lean_Parser_Termination_decreasingBy___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_Termination_decreasingBy___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; -x_1 = l_Lean_Parser_Termination_decreasingBy___closed__1; -x_2 = l_Lean_Parser_Termination_decreasingBy___closed__2; -x_3 = 1; -x_4 = 0; -x_5 = l_Lean_Parser_mkAntiquot(x_1, x_2, x_3, x_4); -return x_5; -} -} -static lean_object* _init_l_Lean_Parser_Termination_decreasingBy___closed__4() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("decreasing_by ", 14); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Termination_decreasingBy___closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Termination_decreasingBy___closed__4; -x_2 = l_Lean_Parser_symbol(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Parser_Termination_decreasingBy___closed__6() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Termination_decreasingBy___closed__5; -x_2 = l_Lean_Parser_Tactic_tacticSeqIndentGt; -x_3 = l_Lean_Parser_andthen(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser_Termination_decreasingBy___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_Termination_decreasingBy___closed__6; -x_3 = l_Lean_Parser_andthen(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser_Termination_decreasingBy___closed__8() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Termination_decreasingBy___closed__2; -x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Termination_decreasingBy___closed__7; -x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); -return x_4; -} -} -static lean_object* _init_l_Lean_Parser_Termination_decreasingBy___closed__9() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Termination_decreasingBy___closed__3; -x_2 = l_Lean_Parser_Termination_decreasingBy___closed__8; -x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser_Termination_decreasingBy___closed__10() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Termination_decreasingBy___closed__2; -x_2 = l_Lean_Parser_Termination_decreasingBy___closed__9; -x_3 = l_Lean_Parser_withCache(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser_Termination_decreasingBy() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Parser_Termination_decreasingBy___closed__10; -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Termination_suffix___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("suffix", 6); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Termination_suffix___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_Command_docComment___closed__1; -x_2 = l_Lean_Parser_Command_docComment___closed__2; -x_3 = l_Lean_Parser_Termination_terminationBy___closed__1; -x_4 = l_Lean_Parser_Termination_suffix___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_Termination_suffix___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; -x_1 = l_Lean_Parser_Termination_suffix___closed__1; -x_2 = l_Lean_Parser_Termination_suffix___closed__2; -x_3 = 1; -x_4 = 0; -x_5 = l_Lean_Parser_mkAntiquot(x_1, x_2, x_3, x_4); -return x_5; -} -} -static lean_object* _init_l_Lean_Parser_Termination_suffix___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Termination_terminationBy_x3f; -x_2 = l_Lean_Parser_Termination_terminationBy; -x_3 = l_Lean_Parser_orelse(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser_Termination_suffix___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_Termination_suffix___closed__4; -x_3 = l_Lean_Parser_andthen(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser_Termination_suffix___closed__6() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Termination_suffix___closed__5; -x_2 = l_Lean_Parser_optional(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Parser_Termination_suffix___closed__7() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Termination_decreasingBy; -x_2 = l_Lean_Parser_optional(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Parser_Termination_suffix___closed__8() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Termination_suffix___closed__6; -x_2 = l_Lean_Parser_Termination_suffix___closed__7; -x_3 = l_Lean_Parser_andthen(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser_Termination_suffix___closed__9() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Termination_suffix___closed__2; -x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Termination_suffix___closed__8; -x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); -return x_4; -} -} -static lean_object* _init_l_Lean_Parser_Termination_suffix___closed__10() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Termination_suffix___closed__3; -x_2 = l_Lean_Parser_Termination_suffix___closed__9; -x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser_Termination_suffix___closed__11() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Termination_suffix___closed__2; -x_2 = l_Lean_Parser_Termination_suffix___closed__10; -x_3 = l_Lean_Parser_withCache(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser_Termination_suffix() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Parser_Termination_suffix___closed__11; -return x_1; -} -} static lean_object* _init_l_Lean_Parser_Term_byTactic___closed__1() { _start: { @@ -8118,7 +7598,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_byTactic_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(155u); +x_1 = lean_unsigned_to_nat(106u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -8130,7 +7610,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_byTactic_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(156u); +x_1 = lean_unsigned_to_nat(107u); x_2 = lean_unsigned_to_nat(55u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -8158,7 +7638,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_byTactic_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(155u); +x_1 = lean_unsigned_to_nat(106u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -8170,7 +7650,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_byTactic_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(155u); +x_1 = lean_unsigned_to_nat(106u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9387,7 +8867,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_ident_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(172u); +x_1 = lean_unsigned_to_nat(123u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9399,7 +8879,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_ident_declRange___clos _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(124u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9427,7 +8907,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_ident_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(172u); +x_1 = lean_unsigned_to_nat(123u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9439,7 +8919,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_ident_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(172u); +x_1 = lean_unsigned_to_nat(123u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9594,7 +9074,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_num_declRange___closed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(174u); +x_1 = lean_unsigned_to_nat(125u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9606,7 +9086,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_num_declRange___closed _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(126u); x_2 = lean_unsigned_to_nat(29u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9634,7 +9114,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_num_declRange___closed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(174u); +x_1 = lean_unsigned_to_nat(125u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9646,7 +9126,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_num_declRange___closed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(174u); +x_1 = lean_unsigned_to_nat(125u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9783,7 +9263,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_scientific_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(176u); +x_1 = lean_unsigned_to_nat(127u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9795,7 +9275,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_scientific_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(177u); +x_1 = lean_unsigned_to_nat(128u); x_2 = lean_unsigned_to_nat(36u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9823,7 +9303,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_scientific_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(176u); +x_1 = lean_unsigned_to_nat(127u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9835,7 +9315,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_scientific_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(176u); +x_1 = lean_unsigned_to_nat(127u); x_2 = lean_unsigned_to_nat(37u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9972,7 +9452,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_str_declRange___closed _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(129u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9984,7 +9464,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_str_declRange___closed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(179u); +x_1 = lean_unsigned_to_nat(130u); x_2 = lean_unsigned_to_nat(29u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10012,7 +9492,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_str_declRange___closed _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(129u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10024,7 +9504,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_str_declRange___closed _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(129u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10161,7 +9641,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_char_declRange___close _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(131u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10173,7 +9653,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_char_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(181u); +x_1 = lean_unsigned_to_nat(132u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10201,7 +9681,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_char_declRange___close _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(131u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10213,7 +9693,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_char_declRange___close _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(131u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10538,7 +10018,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_type_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(183u); +x_1 = lean_unsigned_to_nat(134u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10550,7 +10030,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_type_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(184u); +x_1 = lean_unsigned_to_nat(135u); x_2 = lean_unsigned_to_nat(98u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10578,7 +10058,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_type_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(183u); +x_1 = lean_unsigned_to_nat(134u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10590,7 +10070,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_type_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(183u); +x_1 = lean_unsigned_to_nat(134u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11099,7 +10579,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_sort_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(186u); +x_1 = lean_unsigned_to_nat(137u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11111,7 +10591,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_sort_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(187u); +x_1 = lean_unsigned_to_nat(138u); x_2 = lean_unsigned_to_nat(98u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11139,7 +10619,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_sort_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(186u); +x_1 = lean_unsigned_to_nat(137u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11151,7 +10631,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_sort_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(186u); +x_1 = lean_unsigned_to_nat(137u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11514,7 +10994,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_prop_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(189u); +x_1 = lean_unsigned_to_nat(140u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11526,7 +11006,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_prop_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(190u); +x_1 = lean_unsigned_to_nat(141u); x_2 = lean_unsigned_to_nat(8u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11554,7 +11034,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_prop_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(189u); +x_1 = lean_unsigned_to_nat(140u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11566,7 +11046,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_prop_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(189u); +x_1 = lean_unsigned_to_nat(140u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11817,30 +11297,47 @@ return x_5; static lean_object* _init_l_Lean_Parser_Term_hole___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_Term_hole___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Term_hole___closed__4; +x_2 = l_Lean_Parser_symbol(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Term_hole___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_hole___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Termination_terminationBy___closed__8; +x_3 = l_Lean_Parser_Term_hole___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_hole___closed__5() { +static lean_object* _init_l_Lean_Parser_Term_hole___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_hole___closed__3; -x_2 = l_Lean_Parser_Term_hole___closed__4; +x_2 = l_Lean_Parser_Term_hole___closed__6; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_hole___closed__6() { +static lean_object* _init_l_Lean_Parser_Term_hole___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_hole___closed__2; -x_2 = l_Lean_Parser_Term_hole___closed__5; +x_2 = l_Lean_Parser_Term_hole___closed__7; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -11849,7 +11346,7 @@ static lean_object* _init_l_Lean_Parser_Term_hole() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Term_hole___closed__6; +x_1 = l_Lean_Parser_Term_hole___closed__8; return x_1; } } @@ -11888,7 +11385,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_hole_declRange___close _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(143u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11900,7 +11397,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_hole_declRange___close _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(144u); x_2 = lean_unsigned_to_nat(5u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11928,7 +11425,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_hole_declRange___close _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(143u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11940,7 +11437,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_hole_declRange___close _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(143u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -12008,7 +11505,7 @@ static lean_object* _init_l_Lean_Parser_Term_hole_formatter___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Termination_terminationBy___closed__7; +x_1 = l_Lean_Parser_Term_hole___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; @@ -12093,7 +11590,7 @@ static lean_object* _init_l_Lean_Parser_Term_hole_parenthesizer___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Termination_terminationBy___closed__7; +x_1 = l_Lean_Parser_Term_hole___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; @@ -12299,7 +11796,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_syntheticHole_declRang _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(147u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -12311,7 +11808,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_syntheticHole_declRang _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(148u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -12339,7 +11836,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_syntheticHole_declRang _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(147u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -12351,7 +11848,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_syntheticHole_declRang _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(147u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -12754,7 +12251,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_omission_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(203u); +x_1 = lean_unsigned_to_nat(154u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -12766,7 +12263,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_omission_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(155u); x_2 = lean_unsigned_to_nat(5u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -12794,7 +12291,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_omission_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(203u); +x_1 = lean_unsigned_to_nat(154u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -12806,7 +12303,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_omission_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(203u); +x_1 = lean_unsigned_to_nat(154u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13145,7 +12642,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(207u); +x_1 = lean_unsigned_to_nat(158u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13157,7 +12654,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(208u); +x_1 = lean_unsigned_to_nat(159u); x_2 = lean_unsigned_to_nat(9u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13185,7 +12682,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(207u); +x_1 = lean_unsigned_to_nat(158u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13197,7 +12694,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(207u); +x_1 = lean_unsigned_to_nat(158u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13563,7 +13060,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(213u); +x_1 = lean_unsigned_to_nat(164u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13575,7 +13072,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(214u); +x_1 = lean_unsigned_to_nat(165u); x_2 = lean_unsigned_to_nat(20u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13603,7 +13100,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(213u); +x_1 = lean_unsigned_to_nat(164u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13615,7 +13112,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(213u); +x_1 = lean_unsigned_to_nat(164u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13927,60 +13424,69 @@ return x_2; static lean_object* _init_l_Lean_Parser_Term_typeAscription___closed__6() { _start: { +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = l_Lean_Parser_termParser(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Term_typeAscription___closed__7() { +_start: +{ lean_object* x_1; x_1 = lean_mk_string_from_bytes(" :", 2); return x_1; } } -static lean_object* _init_l_Lean_Parser_Term_typeAscription___closed__7() { +static lean_object* _init_l_Lean_Parser_Term_typeAscription___closed__8() { _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__7; x_2 = l_Lean_Parser_symbol(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_typeAscription___closed__8() { +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_Termination_terminationBy___closed__15; +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__9() { +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__8; +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__7; -x_2 = l_Lean_Parser_Term_typeAscription___closed__9; +x_1 = l_Lean_Parser_Term_typeAscription___closed__8; +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_Termination_terminationBy___closed__15; -x_2 = l_Lean_Parser_Term_typeAscription___closed__10; +x_1 = l_Lean_Parser_Term_typeAscription___closed__6; +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; @@ -13988,17 +13494,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; @@ -14006,17 +13512,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; @@ -14024,62 +13530,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; } @@ -14088,7 +13594,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; } } @@ -14127,7 +13633,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(221u); +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); @@ -14139,7 +13645,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(222u); +x_1 = lean_unsigned_to_nat(173u); x_2 = lean_unsigned_to_nat(109u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -14167,7 +13673,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(221u); +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); @@ -14179,7 +13685,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(221u); +x_1 = lean_unsigned_to_nat(172u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -14257,7 +13763,7 @@ static lean_object* _init_l_Lean_Parser_Term_typeAscription_formatter___closed__ _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__7; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -14349,7 +13855,7 @@ static lean_object* _init_l_Lean_Parser_Term_typeAscription_formatter___closed__ _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; @@ -14478,7 +13984,7 @@ static lean_object* _init_l_Lean_Parser_Term_typeAscription_parenthesizer___clos _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__7; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -14562,7 +14068,7 @@ static lean_object* _init_l_Lean_Parser_Term_typeAscription_parenthesizer___clos _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; @@ -14702,7 +14208,7 @@ static lean_object* _init_l_Lean_Parser_Term_tuple___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_Termination_terminationBy___closed__15; +x_1 = l_Lean_Parser_Term_typeAscription___closed__6; x_2 = l_Lean_Parser_Term_tuple___closed__4; x_3 = l_Lean_Parser_Term_tuple___closed__5; x_4 = 1; @@ -14724,7 +14230,7 @@ static lean_object* _init_l_Lean_Parser_Term_tuple___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Termination_terminationBy___closed__15; +x_1 = l_Lean_Parser_Term_typeAscription___closed__6; x_2 = l_Lean_Parser_Term_tuple___closed__7; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -14734,7 +14240,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; @@ -14744,7 +14250,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; @@ -14764,7 +14270,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; } @@ -14853,7 +14359,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(224u); +x_1 = lean_unsigned_to_nat(175u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -14865,7 +14371,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(225u); +x_1 = lean_unsigned_to_nat(176u); x_2 = lean_unsigned_to_nat(135u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -14893,7 +14399,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(224u); +x_1 = lean_unsigned_to_nat(175u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -14905,7 +14411,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(224u); +x_1 = lean_unsigned_to_nat(175u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -15347,8 +14853,8 @@ 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_2 = l_Lean_Parser_Termination_terminationBy___closed__15; +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; } @@ -15357,7 +14863,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; @@ -15368,7 +14874,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; } @@ -15457,7 +14963,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(235u); +x_1 = lean_unsigned_to_nat(186u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -15469,7 +14975,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(236u); +x_1 = lean_unsigned_to_nat(187u); x_2 = lean_unsigned_to_nat(81u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -15497,7 +15003,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(235u); +x_1 = lean_unsigned_to_nat(186u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -15509,7 +15015,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(235u); +x_1 = lean_unsigned_to_nat(186u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -15866,7 +15372,7 @@ static lean_object* _init_l_Lean_Parser_Term_anonymousCtor___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_Termination_terminationBy___closed__15; +x_1 = l_Lean_Parser_Term_typeAscription___closed__6; x_2 = l_Lean_Parser_Term_tuple___closed__4; x_3 = l_Lean_Parser_Term_tuple___closed__5; x_4 = 1; @@ -15878,7 +15384,7 @@ 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__12; +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; @@ -15888,7 +15394,7 @@ 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__14; +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; @@ -16005,7 +15511,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(244u); +x_1 = lean_unsigned_to_nat(195u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -16017,7 +15523,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(245u); +x_1 = lean_unsigned_to_nat(196u); x_2 = lean_unsigned_to_nat(101u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -16045,7 +15551,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(244u); +x_1 = lean_unsigned_to_nat(195u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -16057,7 +15563,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(244u); +x_1 = lean_unsigned_to_nat(195u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -16522,7 +16028,7 @@ static lean_object* _init_l_Lean_Parser_Term_fromTerm___closed__6() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_fromTerm___closed__5; -x_2 = l_Lean_Parser_Termination_terminationBy___closed__15; +x_2 = l_Lean_Parser_Term_typeAscription___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -16687,7 +16193,7 @@ 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_Termination_terminationBy___closed__15; +x_1 = l_Lean_Parser_Term_typeAscription___closed__6; x_2 = l_Lean_Parser_Term_sufficesDecl___closed__10; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -16814,7 +16320,7 @@ static lean_object* _init_l_Lean_Parser_Term_suffices___closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Termination_terminationBy___closed__15; +x_1 = l_Lean_Parser_Term_typeAscription___closed__6; x_2 = l_Lean_Parser_Term_optSemicolon(x_1); return x_2; } @@ -16885,7 +16391,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(255u); +x_1 = lean_unsigned_to_nat(206u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -16897,7 +16403,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(256u); +x_1 = lean_unsigned_to_nat(207u); x_2 = lean_unsigned_to_nat(71u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -16925,7 +16431,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(255u); +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); @@ -16937,7 +16443,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(255u); +x_1 = lean_unsigned_to_nat(206u); x_2 = lean_unsigned_to_nat(37u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -18259,7 +17765,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(257u); +x_1 = lean_unsigned_to_nat(208u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -18271,7 +17777,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(257u); +x_1 = lean_unsigned_to_nat(208u); x_2 = lean_unsigned_to_nat(108u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -18299,7 +17805,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(257u); +x_1 = lean_unsigned_to_nat(208u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -18311,7 +17817,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(257u); +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); @@ -18604,8 +18110,8 @@ 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_2 = l_Lean_Parser_Termination_terminationBy___closed__15; +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; } @@ -18890,7 +18396,7 @@ static lean_object* _init_l_Lean_Parser_Term_structInstField___closed__6() { { 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_Termination_terminationBy___closed__15; +x_2 = l_Lean_Parser_Term_typeAscription___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -19239,7 +18745,7 @@ static lean_object* _init_l_Lean_Parser_Term_structInst___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_Termination_terminationBy___closed__15; +x_1 = l_Lean_Parser_Term_typeAscription___closed__6; x_2 = l_Lean_Parser_Term_tuple___closed__4; x_3 = l_Lean_Parser_Term_tuple___closed__5; x_4 = 0; @@ -19359,7 +18865,7 @@ static lean_object* _init_l_Lean_Parser_Term_structInst___closed__18() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_optIdent___closed__2; -x_2 = l_Lean_Parser_Termination_terminationBy___closed__15; +x_2 = l_Lean_Parser_Term_typeAscription___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -19407,7 +18913,7 @@ static lean_object* _init_l_Lean_Parser_Term_structInst___closed__23() { _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__22; x_3 = l_Lean_Parser_adaptCacheableContext(x_1, x_2); return x_3; @@ -19524,7 +19030,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(279u); +x_1 = lean_unsigned_to_nat(230u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19536,7 +19042,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(283u); +x_1 = lean_unsigned_to_nat(234u); x_2 = lean_unsigned_to_nat(46u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19564,7 +19070,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(279u); +x_1 = lean_unsigned_to_nat(230u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19576,7 +19082,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(279u); +x_1 = lean_unsigned_to_nat(230u); x_2 = lean_unsigned_to_nat(37u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -21525,7 +21031,7 @@ static lean_object* _init_l_Lean_Parser_Term_typeSpec___closed__4() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_optIdent___closed__2; -x_2 = l_Lean_Parser_Termination_terminationBy___closed__15; +x_2 = l_Lean_Parser_Term_typeAscription___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -21728,7 +21234,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(290u); +x_1 = lean_unsigned_to_nat(241u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -21740,7 +21246,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(291u); +x_1 = lean_unsigned_to_nat(242u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -21768,7 +21274,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(290u); +x_1 = lean_unsigned_to_nat(241u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -21780,7 +21286,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(290u); +x_1 = lean_unsigned_to_nat(241u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22084,7 +21590,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; } @@ -22173,7 +21679,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(296u); +x_1 = lean_unsigned_to_nat(247u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22185,7 +21691,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(297u); +x_1 = lean_unsigned_to_nat(248u); x_2 = lean_unsigned_to_nat(43u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22213,7 +21719,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(296u); +x_1 = lean_unsigned_to_nat(247u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22225,7 +21731,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(296u); +x_1 = lean_unsigned_to_nat(247u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22929,9 +22435,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); @@ -22992,7 +22498,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); @@ -23240,7 +22746,7 @@ static lean_object* _init_l_Lean_Parser_Term_instBinder___closed__4() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_optIdent; -x_2 = l_Lean_Parser_Termination_terminationBy___closed__15; +x_2 = l_Lean_Parser_Term_typeAscription___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -23249,7 +22755,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; @@ -23443,7 +22949,7 @@ static lean_object* _init_l_Lean_Parser_Term_depArrow___closed__8() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_depArrow___closed__7; -x_2 = l_Lean_Parser_Termination_terminationBy___closed__15; +x_2 = l_Lean_Parser_Term_typeAscription___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -23514,7 +23020,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_depArrow_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(365u); +x_1 = lean_unsigned_to_nat(316u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23526,7 +23032,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_depArrow_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(366u); +x_1 = lean_unsigned_to_nat(317u); x_2 = lean_unsigned_to_nat(66u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23554,7 +23060,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_depArrow_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(365u); +x_1 = lean_unsigned_to_nat(316u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23566,7 +23072,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_depArrow_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(365u); +x_1 = lean_unsigned_to_nat(316u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -25604,7 +25110,7 @@ static lean_object* _init_l_Lean_Parser_Term_forall___closed__10() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_tuple___closed__5; -x_2 = l_Lean_Parser_Termination_terminationBy___closed__15; +x_2 = l_Lean_Parser_Term_typeAscription___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -25695,7 +25201,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_forall_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(369u); +x_1 = lean_unsigned_to_nat(320u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -25707,7 +25213,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_forall_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(372u); +x_1 = lean_unsigned_to_nat(323u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -25735,7 +25241,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_forall_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(369u); +x_1 = lean_unsigned_to_nat(320u); x_2 = lean_unsigned_to_nat(4u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -25747,7 +25253,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_forall_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(369u); +x_1 = lean_unsigned_to_nat(320u); x_2 = lean_unsigned_to_nat(12u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -26455,7 +25961,7 @@ static lean_object* _init_l_Lean_Parser_Term_matchAltExpr___closed__1() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Termination_terminationBy___closed__15; +x_1 = l_Lean_Parser_Term_typeAscription___closed__6; x_2 = l_Lean_Parser_Term_matchAlt(x_1); return x_2; } @@ -26573,7 +26079,7 @@ static lean_object* _init_l_Lean_Parser_Term_matchDiscr___closed__4() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_optIdent___closed__5; -x_2 = l_Lean_Parser_Termination_terminationBy___closed__15; +x_2 = l_Lean_Parser_Term_typeAscription___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -26878,7 +26384,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; @@ -27183,7 +26689,7 @@ 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_Termination_terminationBy___closed__15; +x_1 = l_Lean_Parser_Term_typeAscription___closed__6; x_2 = l_Lean_Parser_Term_matchAlts(x_1); return x_2; } @@ -27312,7 +26818,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_match_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(439u); +x_1 = lean_unsigned_to_nat(390u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27324,7 +26830,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_match_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(441u); +x_1 = lean_unsigned_to_nat(392u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27352,7 +26858,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_match_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(439u); +x_1 = lean_unsigned_to_nat(390u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27364,7 +26870,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_match_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(439u); +x_1 = lean_unsigned_to_nat(390u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29455,7 +28961,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_nomatch_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(447u); +x_1 = lean_unsigned_to_nat(398u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29467,7 +28973,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_nomatch_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(447u); +x_1 = lean_unsigned_to_nat(398u); x_2 = lean_unsigned_to_nat(100u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29495,7 +29001,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_nomatch_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(447u); +x_1 = lean_unsigned_to_nat(398u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29507,7 +29013,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_nomatch_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(447u); +x_1 = lean_unsigned_to_nat(398u); x_2 = lean_unsigned_to_nat(36u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29844,7 +29350,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_nofun_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(449u); +x_1 = lean_unsigned_to_nat(400u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29856,7 +29362,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_nofun_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(449u); +x_1 = lean_unsigned_to_nat(400u); x_2 = lean_unsigned_to_nat(60u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29884,7 +29390,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_nofun_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(449u); +x_1 = lean_unsigned_to_nat(400u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29896,7 +29402,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_nofun_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(449u); +x_1 = lean_unsigned_to_nat(400u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30452,7 +29958,7 @@ static lean_object* _init_l_Lean_Parser_Term_basicFun___closed__11() { { 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_typeAscription___closed__8; +x_2 = l_Lean_Parser_Term_typeAscription___closed__9; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -30632,7 +30138,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_fun_declRange___closed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(465u); +x_1 = lean_unsigned_to_nat(416u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30644,7 +30150,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_fun_declRange___closed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(466u); +x_1 = lean_unsigned_to_nat(417u); x_2 = lean_unsigned_to_nat(73u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30672,7 +30178,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_fun_declRange___closed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(465u); +x_1 = lean_unsigned_to_nat(416u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30684,7 +30190,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_fun_declRange___closed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(465u); +x_1 = lean_unsigned_to_nat(416u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -31936,7 +31442,7 @@ static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot___closed__10( { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_generalizingParam___closed__8; -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; } @@ -32045,7 +31551,7 @@ static lean_object* _init_l_Lean_Parser_Term_leading__parser___closed__6() { { 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_typeAscription___closed__8; +x_2 = l_Lean_Parser_Term_typeAscription___closed__9; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -32126,7 +31632,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_leading__parser_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(472u); +x_1 = lean_unsigned_to_nat(423u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -32138,7 +31644,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_leading__parser_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(473u); +x_1 = lean_unsigned_to_nat(424u); x_2 = lean_unsigned_to_nat(98u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -32166,7 +31672,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_leading__parser_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(472u); +x_1 = lean_unsigned_to_nat(423u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -32178,7 +31684,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_leading__parser_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(472u); +x_1 = lean_unsigned_to_nat(423u); x_2 = lean_unsigned_to_nat(43u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -32942,7 +32448,7 @@ static lean_object* _init_l_Lean_Parser_Term_trailing__parser___closed__5() { { 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__8; +x_2 = l_Lean_Parser_Term_typeAscription___closed__9; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -33023,7 +32529,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_trailing__parser_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(474u); +x_1 = lean_unsigned_to_nat(425u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -33035,7 +32541,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_trailing__parser_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(475u); +x_1 = lean_unsigned_to_nat(426u); x_2 = lean_unsigned_to_nat(86u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -33063,7 +32569,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_trailing__parser_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(474u); +x_1 = lean_unsigned_to_nat(425u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -33075,7 +32581,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_trailing__parser_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(474u); +x_1 = lean_unsigned_to_nat(425u); x_2 = lean_unsigned_to_nat(44u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -33487,7 +32993,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_borrowed_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(477u); +x_1 = lean_unsigned_to_nat(428u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -33499,7 +33005,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_borrowed_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(478u); +x_1 = lean_unsigned_to_nat(429u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -33527,7 +33033,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_borrowed_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(477u); +x_1 = lean_unsigned_to_nat(428u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -33539,7 +33045,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_borrowed_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(477u); +x_1 = lean_unsigned_to_nat(428u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -33895,7 +33401,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_quotedName_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(480u); +x_1 = lean_unsigned_to_nat(431u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -33907,7 +33413,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_quotedName_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(480u); +x_1 = lean_unsigned_to_nat(431u); x_2 = lean_unsigned_to_nat(63u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -33935,7 +33441,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_quotedName_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(480u); +x_1 = lean_unsigned_to_nat(431u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -33947,7 +33453,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_quotedName_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(480u); +x_1 = lean_unsigned_to_nat(431u); x_2 = lean_unsigned_to_nat(37u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -34339,7 +33845,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_doubleQuotedName_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(487u); +x_1 = lean_unsigned_to_nat(438u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -34351,7 +33857,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_doubleQuotedName_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(488u); +x_1 = lean_unsigned_to_nat(439u); x_2 = lean_unsigned_to_nat(68u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -34379,7 +33885,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_doubleQuotedName_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(487u); +x_1 = lean_unsigned_to_nat(438u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -34391,7 +33897,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_doubleQuotedName_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(487u); +x_1 = lean_unsigned_to_nat(438u); x_2 = lean_unsigned_to_nat(43u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -34931,7 +34437,7 @@ static lean_object* _init_l_Lean_Parser_Term_letIdDecl___closed__6() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_letIdDecl___closed__5; -x_2 = l_Lean_Parser_Termination_terminationBy___closed__15; +x_2 = l_Lean_Parser_Term_typeAscription___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -35030,7 +34536,7 @@ static lean_object* _init_l_Lean_Parser_Term_letPatDecl___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Termination_terminationBy___closed__15; +x_1 = l_Lean_Parser_Term_typeAscription___closed__6; x_2 = l_Lean_Parser_Term_letPatDecl___closed__5; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -35050,7 +34556,7 @@ static lean_object* _init_l_Lean_Parser_Term_letPatDecl___closed__8() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_letPatDecl___closed__7; -x_2 = l_Lean_Parser_Termination_terminationBy___closed__15; +x_2 = l_Lean_Parser_Term_typeAscription___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -35464,7 +34970,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_let_declRange___closed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(550u); +x_1 = lean_unsigned_to_nat(501u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -35476,7 +34982,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_let_declRange___closed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(551u); +x_1 = lean_unsigned_to_nat(502u); x_2 = lean_unsigned_to_nat(61u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -35504,7 +35010,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_let_declRange___closed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(550u); +x_1 = lean_unsigned_to_nat(501u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -35516,7 +35022,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_let_declRange___closed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(550u); +x_1 = lean_unsigned_to_nat(501u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -37163,7 +36669,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_let__fun_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(557u); +x_1 = lean_unsigned_to_nat(508u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -37175,7 +36681,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_let__fun_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(558u); +x_1 = lean_unsigned_to_nat(509u); x_2 = lean_unsigned_to_nat(87u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -37203,7 +36709,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_let__fun_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(557u); +x_1 = lean_unsigned_to_nat(508u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -37215,7 +36721,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_let__fun_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(557u); +x_1 = lean_unsigned_to_nat(508u); x_2 = lean_unsigned_to_nat(36u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -37695,7 +37201,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_let__delayed_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(562u); +x_1 = lean_unsigned_to_nat(513u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -37707,7 +37213,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_let__delayed_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(563u); +x_1 = lean_unsigned_to_nat(514u); x_2 = lean_unsigned_to_nat(69u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -37735,7 +37241,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_let__delayed_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(562u); +x_1 = lean_unsigned_to_nat(513u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -37747,7 +37253,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_let__delayed_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(562u); +x_1 = lean_unsigned_to_nat(513u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -38183,7 +37689,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_let__tmp_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(568u); +x_1 = lean_unsigned_to_nat(519u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -38195,7 +37701,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_let__tmp_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(569u); +x_1 = lean_unsigned_to_nat(520u); x_2 = lean_unsigned_to_nat(65u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -38223,7 +37729,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_let__tmp_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(568u); +x_1 = lean_unsigned_to_nat(519u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -38235,7 +37741,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_let__tmp_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(568u); +x_1 = lean_unsigned_to_nat(519u); x_2 = lean_unsigned_to_nat(36u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -38612,7 +38118,7 @@ static lean_object* _init_l_Lean_Parser_Term_haveIdDecl___closed__6() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_haveIdDecl___closed__5; -x_2 = l_Lean_Parser_Termination_terminationBy___closed__15; +x_2 = l_Lean_Parser_Term_typeAscription___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -38962,7 +38468,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_have_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(582u); +x_1 = lean_unsigned_to_nat(533u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -38974,7 +38480,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_have_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(583u); +x_1 = lean_unsigned_to_nat(534u); x_2 = lean_unsigned_to_nat(62u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -39002,7 +38508,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_have_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(582u); +x_1 = lean_unsigned_to_nat(533u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -39014,7 +38520,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_have_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(582u); +x_1 = lean_unsigned_to_nat(533u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -40142,7 +39648,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_haveI_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(585u); +x_1 = lean_unsigned_to_nat(536u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -40154,7 +39660,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_haveI_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(586u); +x_1 = lean_unsigned_to_nat(537u); x_2 = lean_unsigned_to_nat(64u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -40182,7 +39688,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_haveI_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(585u); +x_1 = lean_unsigned_to_nat(536u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -40194,7 +39700,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_haveI_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(585u); +x_1 = lean_unsigned_to_nat(536u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -40630,7 +40136,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_letI_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(588u); +x_1 = lean_unsigned_to_nat(539u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -40642,7 +40148,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_letI_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(589u); +x_1 = lean_unsigned_to_nat(540u); x_2 = lean_unsigned_to_nat(63u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -40670,7 +40176,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_letI_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(588u); +x_1 = lean_unsigned_to_nat(539u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -40682,7 +40188,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_letI_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(588u); +x_1 = lean_unsigned_to_nat(539u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -41406,7 +40912,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; @@ -41488,6 +40994,474 @@ x_1 = l_Lean_Parser_Term_attributes___closed__14; return x_1; } } +static lean_object* _init_l_Lean_Parser_Termination_terminationBy___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Termination", 11); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Termination_terminationBy___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("terminationBy", 13); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Termination_terminationBy___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_Command_docComment___closed__1; +x_2 = l_Lean_Parser_Command_docComment___closed__2; +x_3 = l_Lean_Parser_Termination_terminationBy___closed__1; +x_4 = l_Lean_Parser_Termination_terminationBy___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_Termination_terminationBy___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; +x_1 = l_Lean_Parser_Termination_terminationBy___closed__2; +x_2 = l_Lean_Parser_Termination_terminationBy___closed__3; +x_3 = 1; +x_4 = 0; +x_5 = l_Lean_Parser_mkAntiquot(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Parser_Termination_terminationBy___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("termination_by ", 15); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Termination_terminationBy___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Termination_terminationBy___closed__5; +x_2 = l_Lean_Parser_symbol(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Termination_terminationBy___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Term_haveIdLhs___closed__1; +x_2 = l_Lean_Parser_many(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Termination_terminationBy___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Termination_terminationBy___closed__7; +x_2 = l_Lean_Parser_darrow___closed__2; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Termination_terminationBy___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Termination_terminationBy___closed__8; +x_2 = l_Lean_Parser_atomic(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Termination_terminationBy___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Termination_terminationBy___closed__9; +x_2 = l_Lean_Parser_optional(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Termination_terminationBy___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Termination_terminationBy___closed__10; +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_Termination_terminationBy___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Termination_terminationBy___closed__6; +x_2 = l_Lean_Parser_Termination_terminationBy___closed__11; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Termination_terminationBy___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Termination_terminationBy___closed__3; +x_2 = lean_unsigned_to_nat(1024u); +x_3 = l_Lean_Parser_Termination_terminationBy___closed__12; +x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Termination_terminationBy___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Termination_terminationBy___closed__4; +x_2 = l_Lean_Parser_Termination_terminationBy___closed__13; +x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Termination_terminationBy___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Termination_terminationBy___closed__3; +x_2 = l_Lean_Parser_Termination_terminationBy___closed__14; +x_3 = l_Lean_Parser_withCache(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Termination_terminationBy() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Parser_Termination_terminationBy___closed__15; +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Termination_terminationBy_x3f___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("terminationBy\?", 14); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Termination_terminationBy_x3f___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_Command_docComment___closed__1; +x_2 = l_Lean_Parser_Command_docComment___closed__2; +x_3 = l_Lean_Parser_Termination_terminationBy___closed__1; +x_4 = l_Lean_Parser_Termination_terminationBy_x3f___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_Termination_terminationBy_x3f___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; +x_1 = l_Lean_Parser_Termination_terminationBy_x3f___closed__1; +x_2 = l_Lean_Parser_Termination_terminationBy_x3f___closed__2; +x_3 = 1; +x_4 = 0; +x_5 = l_Lean_Parser_mkAntiquot(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Parser_Termination_terminationBy_x3f___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("termination_by\?", 15); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Termination_terminationBy_x3f___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Termination_terminationBy_x3f___closed__4; +x_2 = l_Lean_Parser_symbol(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Termination_terminationBy_x3f___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Termination_terminationBy_x3f___closed__2; +x_2 = lean_unsigned_to_nat(1024u); +x_3 = l_Lean_Parser_Termination_terminationBy_x3f___closed__5; +x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Termination_terminationBy_x3f___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Termination_terminationBy_x3f___closed__3; +x_2 = l_Lean_Parser_Termination_terminationBy_x3f___closed__6; +x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Termination_terminationBy_x3f___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Termination_terminationBy_x3f___closed__2; +x_2 = l_Lean_Parser_Termination_terminationBy_x3f___closed__7; +x_3 = l_Lean_Parser_withCache(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Termination_terminationBy_x3f() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Parser_Termination_terminationBy_x3f___closed__8; +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Termination_decreasingBy___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("decreasingBy", 12); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Termination_decreasingBy___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_Command_docComment___closed__1; +x_2 = l_Lean_Parser_Command_docComment___closed__2; +x_3 = l_Lean_Parser_Termination_terminationBy___closed__1; +x_4 = l_Lean_Parser_Termination_decreasingBy___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_Termination_decreasingBy___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; +x_1 = l_Lean_Parser_Termination_decreasingBy___closed__1; +x_2 = l_Lean_Parser_Termination_decreasingBy___closed__2; +x_3 = 1; +x_4 = 0; +x_5 = l_Lean_Parser_mkAntiquot(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Parser_Termination_decreasingBy___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("decreasing_by ", 14); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Termination_decreasingBy___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Termination_decreasingBy___closed__4; +x_2 = l_Lean_Parser_symbol(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Termination_decreasingBy___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Termination_decreasingBy___closed__5; +x_2 = l_Lean_Parser_Tactic_tacticSeqIndentGt; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Termination_decreasingBy___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_Termination_decreasingBy___closed__6; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Termination_decreasingBy___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Termination_decreasingBy___closed__2; +x_2 = lean_unsigned_to_nat(1024u); +x_3 = l_Lean_Parser_Termination_decreasingBy___closed__7; +x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Termination_decreasingBy___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Termination_decreasingBy___closed__3; +x_2 = l_Lean_Parser_Termination_decreasingBy___closed__8; +x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Termination_decreasingBy___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Termination_decreasingBy___closed__2; +x_2 = l_Lean_Parser_Termination_decreasingBy___closed__9; +x_3 = l_Lean_Parser_withCache(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Termination_decreasingBy() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Parser_Termination_decreasingBy___closed__10; +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Termination_suffix___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("suffix", 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Termination_suffix___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_Command_docComment___closed__1; +x_2 = l_Lean_Parser_Command_docComment___closed__2; +x_3 = l_Lean_Parser_Termination_terminationBy___closed__1; +x_4 = l_Lean_Parser_Termination_suffix___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_Termination_suffix___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; +x_1 = l_Lean_Parser_Termination_suffix___closed__1; +x_2 = l_Lean_Parser_Termination_suffix___closed__2; +x_3 = 1; +x_4 = 0; +x_5 = l_Lean_Parser_mkAntiquot(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Parser_Termination_suffix___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Termination_terminationBy_x3f; +x_2 = l_Lean_Parser_Termination_terminationBy; +x_3 = l_Lean_Parser_orelse(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Termination_suffix___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_Termination_suffix___closed__4; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Termination_suffix___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Termination_suffix___closed__5; +x_2 = l_Lean_Parser_optional(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Termination_suffix___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Termination_decreasingBy; +x_2 = l_Lean_Parser_optional(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Termination_suffix___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Termination_suffix___closed__6; +x_2 = l_Lean_Parser_Termination_suffix___closed__7; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Termination_suffix___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Termination_suffix___closed__2; +x_2 = lean_unsigned_to_nat(1024u); +x_3 = l_Lean_Parser_Termination_suffix___closed__8; +x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Termination_suffix___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Termination_suffix___closed__3; +x_2 = l_Lean_Parser_Termination_suffix___closed__9; +x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Termination_suffix___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Termination_suffix___closed__2; +x_2 = l_Lean_Parser_Termination_suffix___closed__10; +x_3 = l_Lean_Parser_withCache(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Termination_suffix() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Parser_Termination_suffix___closed__11; +return x_1; +} +} static lean_object* _init_l_Lean_Parser_Term_letRecDecl___closed__1() { _start: { @@ -41845,7 +41819,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_letrec_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(607u); +x_1 = lean_unsigned_to_nat(611u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -41857,7 +41831,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_letrec_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(609u); +x_1 = lean_unsigned_to_nat(613u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -41885,7 +41859,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_letrec_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(607u); +x_1 = lean_unsigned_to_nat(611u); x_2 = lean_unsigned_to_nat(4u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -41897,7 +41871,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_letrec_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(607u); +x_1 = lean_unsigned_to_nat(611u); x_2 = lean_unsigned_to_nat(12u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -42727,38 +42701,14 @@ return x_2; static lean_object* _init_l_Lean_Parser_Termination_terminationBy_formatter___closed__3() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_ident_formatter___closed__2; -x_2 = l_Lean_Parser_Term_hole_formatter___closed__2; -x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_orelse_formatter), 7, 2); -lean_closure_set(x_3, 0, x_1); -lean_closure_set(x_3, 1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser_Termination_terminationBy_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_2 = l_Lean_Parser_Termination_terminationBy_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_Termination_terminationBy_formatter___closed__5() { -_start: -{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Termination_terminationBy_formatter___closed__4; +x_1 = l_Lean_Parser_Term_haveIdLhs_formatter___closed__1; 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_Termination_terminationBy_formatter___closed__6() { +static lean_object* _init_l_Lean_Parser_Termination_terminationBy_formatter___closed__4() { _start: { lean_object* x_1; lean_object* x_2; @@ -42768,43 +42718,43 @@ lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Termination_terminationBy_formatter___closed__7() { +static lean_object* _init_l_Lean_Parser_Termination_terminationBy_formatter___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Termination_terminationBy_formatter___closed__5; -x_2 = l_Lean_Parser_Termination_terminationBy_formatter___closed__6; +x_1 = l_Lean_Parser_Termination_terminationBy_formatter___closed__3; +x_2 = l_Lean_Parser_Termination_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); lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Termination_terminationBy_formatter___closed__8() { +static lean_object* _init_l_Lean_Parser_Termination_terminationBy_formatter___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Termination_terminationBy_formatter___closed__7; +x_1 = l_Lean_Parser_Termination_terminationBy_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_Termination_terminationBy_formatter___closed__9() { +static lean_object* _init_l_Lean_Parser_Termination_terminationBy_formatter___closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Termination_terminationBy_formatter___closed__8; +x_1 = l_Lean_Parser_Termination_terminationBy_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_Termination_terminationBy_formatter___closed__10() { +static lean_object* _init_l_Lean_Parser_Termination_terminationBy_formatter___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Termination_terminationBy_formatter___closed__9; +x_1 = l_Lean_Parser_Termination_terminationBy_formatter___closed__7; 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); @@ -42812,25 +42762,25 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Termination_terminationBy_formatter___closed__11() { +static lean_object* _init_l_Lean_Parser_Termination_terminationBy_formatter___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Termination_terminationBy_formatter___closed__2; -x_2 = l_Lean_Parser_Termination_terminationBy_formatter___closed__10; +x_2 = l_Lean_Parser_Termination_terminationBy_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_Termination_terminationBy_formatter___closed__12() { +static lean_object* _init_l_Lean_Parser_Termination_terminationBy_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_Termination_terminationBy___closed__3; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Termination_terminationBy_formatter___closed__11; +x_3 = l_Lean_Parser_Termination_terminationBy_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); @@ -42843,7 +42793,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Termination_terminationBy_formatter(lean_ { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Termination_terminationBy_formatter___closed__1; -x_7 = l_Lean_Parser_Termination_terminationBy_formatter___closed__12; +x_7 = l_Lean_Parser_Termination_terminationBy_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; } @@ -44263,38 +44213,14 @@ return x_2; static lean_object* _init_l_Lean_Parser_Termination_terminationBy_parenthesizer___closed__3() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_ident_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Term_hole_parenthesizer___closed__2; -x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_orelse_parenthesizer), 7, 2); -lean_closure_set(x_3, 0, x_1); -lean_closure_set(x_3, 1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser_Termination_terminationBy_parenthesizer___closed__4() { -_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_Termination_terminationBy_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_Termination_terminationBy_parenthesizer___closed__5() { -_start: -{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Termination_terminationBy_parenthesizer___closed__4; +x_1 = l_Lean_Parser_Term_haveIdLhs_parenthesizer___closed__1; 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_Termination_terminationBy_parenthesizer___closed__6() { +static lean_object* _init_l_Lean_Parser_Termination_terminationBy_parenthesizer___closed__4() { _start: { lean_object* x_1; lean_object* x_2; @@ -44304,43 +44230,43 @@ lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Termination_terminationBy_parenthesizer___closed__7() { +static lean_object* _init_l_Lean_Parser_Termination_terminationBy_parenthesizer___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Termination_terminationBy_parenthesizer___closed__5; -x_2 = l_Lean_Parser_Termination_terminationBy_parenthesizer___closed__6; +x_1 = l_Lean_Parser_Termination_terminationBy_parenthesizer___closed__3; +x_2 = l_Lean_Parser_Termination_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); lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Termination_terminationBy_parenthesizer___closed__8() { +static lean_object* _init_l_Lean_Parser_Termination_terminationBy_parenthesizer___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Termination_terminationBy_parenthesizer___closed__7; +x_1 = l_Lean_Parser_Termination_terminationBy_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_Termination_terminationBy_parenthesizer___closed__9() { +static lean_object* _init_l_Lean_Parser_Termination_terminationBy_parenthesizer___closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Termination_terminationBy_parenthesizer___closed__8; +x_1 = l_Lean_Parser_Termination_terminationBy_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_Termination_terminationBy_parenthesizer___closed__10() { +static lean_object* _init_l_Lean_Parser_Termination_terminationBy_parenthesizer___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Termination_terminationBy_parenthesizer___closed__9; +x_1 = l_Lean_Parser_Termination_terminationBy_parenthesizer___closed__7; 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); @@ -44348,25 +44274,25 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Termination_terminationBy_parenthesizer___closed__11() { +static lean_object* _init_l_Lean_Parser_Termination_terminationBy_parenthesizer___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Termination_terminationBy_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Termination_terminationBy_parenthesizer___closed__10; +x_2 = l_Lean_Parser_Termination_terminationBy_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_Termination_terminationBy_parenthesizer___closed__12() { +static lean_object* _init_l_Lean_Parser_Termination_terminationBy_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_Termination_terminationBy___closed__3; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Termination_terminationBy_parenthesizer___closed__11; +x_3 = l_Lean_Parser_Termination_terminationBy_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); @@ -44379,7 +44305,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Termination_terminationBy_parenthesizer(l { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Termination_terminationBy_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Termination_terminationBy_parenthesizer___closed__12; +x_7 = l_Lean_Parser_Termination_terminationBy_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; } @@ -45886,7 +45812,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_noindex_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(619u); +x_1 = lean_unsigned_to_nat(623u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -45898,7 +45824,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_noindex_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(620u); +x_1 = lean_unsigned_to_nat(624u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -45926,7 +45852,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_noindex_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(619u); +x_1 = lean_unsigned_to_nat(623u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -45938,7 +45864,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_noindex_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(619u); +x_1 = lean_unsigned_to_nat(623u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -46232,7 +46158,7 @@ static lean_object* _init_l_Lean_Parser_Term_unsafe___closed__6() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_unsafe___closed__5; -x_2 = l_Lean_Parser_Termination_terminationBy___closed__15; +x_2 = l_Lean_Parser_Term_typeAscription___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -46311,7 +46237,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_unsafe_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(633u); +x_1 = lean_unsigned_to_nat(637u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -46323,7 +46249,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_unsafe_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(633u); +x_1 = lean_unsigned_to_nat(637u); x_2 = lean_unsigned_to_nat(86u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -46351,7 +46277,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_unsafe_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(633u); +x_1 = lean_unsigned_to_nat(637u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -46363,7 +46289,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_unsafe_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(633u); +x_1 = lean_unsigned_to_nat(637u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -46776,7 +46702,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_binrel_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(636u); +x_1 = lean_unsigned_to_nat(640u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -46788,7 +46714,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_binrel_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(637u); +x_1 = lean_unsigned_to_nat(641u); x_2 = lean_unsigned_to_nat(87u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -46816,7 +46742,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_binrel_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(636u); +x_1 = lean_unsigned_to_nat(640u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -46828,7 +46754,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_binrel_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(636u); +x_1 = lean_unsigned_to_nat(640u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -47285,7 +47211,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_binrel__no__prop_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(639u); +x_1 = lean_unsigned_to_nat(643u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -47297,7 +47223,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_binrel__no__prop_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(640u); +x_1 = lean_unsigned_to_nat(644u); x_2 = lean_unsigned_to_nat(95u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -47325,7 +47251,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_binrel__no__prop_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(639u); +x_1 = lean_unsigned_to_nat(643u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -47337,7 +47263,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_binrel__no__prop_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(639u); +x_1 = lean_unsigned_to_nat(643u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -47710,7 +47636,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_binop_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(642u); +x_1 = lean_unsigned_to_nat(646u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -47722,7 +47648,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_binop_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(643u); +x_1 = lean_unsigned_to_nat(647u); x_2 = lean_unsigned_to_nat(86u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -47750,7 +47676,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_binop_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(642u); +x_1 = lean_unsigned_to_nat(646u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -47762,7 +47688,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_binop_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(642u); +x_1 = lean_unsigned_to_nat(646u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -48135,7 +48061,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_binop__lazy_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(645u); +x_1 = lean_unsigned_to_nat(649u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -48147,7 +48073,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_binop__lazy_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(646u); +x_1 = lean_unsigned_to_nat(650u); x_2 = lean_unsigned_to_nat(91u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -48175,7 +48101,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_binop__lazy_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(645u); +x_1 = lean_unsigned_to_nat(649u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -48187,7 +48113,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_binop__lazy_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(645u); +x_1 = lean_unsigned_to_nat(649u); x_2 = lean_unsigned_to_nat(37u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -48560,7 +48486,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_leftact_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(649u); +x_1 = lean_unsigned_to_nat(653u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -48572,7 +48498,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_leftact_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(650u); +x_1 = lean_unsigned_to_nat(654u); x_2 = lean_unsigned_to_nat(88u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -48600,7 +48526,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_leftact_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(649u); +x_1 = lean_unsigned_to_nat(653u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -48612,7 +48538,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_leftact_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(649u); +x_1 = lean_unsigned_to_nat(653u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -48985,7 +48911,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_rightact_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(653u); +x_1 = lean_unsigned_to_nat(657u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -48997,7 +48923,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_rightact_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(654u); +x_1 = lean_unsigned_to_nat(658u); x_2 = lean_unsigned_to_nat(89u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -49025,7 +48951,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_rightact_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(653u); +x_1 = lean_unsigned_to_nat(657u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -49037,7 +48963,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_rightact_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(653u); +x_1 = lean_unsigned_to_nat(657u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -49420,7 +49346,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_unop_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(656u); +x_1 = lean_unsigned_to_nat(660u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -49432,7 +49358,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_unop_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(657u); +x_1 = lean_unsigned_to_nat(661u); x_2 = lean_unsigned_to_nat(52u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -49460,7 +49386,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_unop_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(656u); +x_1 = lean_unsigned_to_nat(660u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -49472,7 +49398,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_unop_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(656u); +x_1 = lean_unsigned_to_nat(660u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -49861,7 +49787,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_forInMacro_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(659u); +x_1 = lean_unsigned_to_nat(663u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -49873,7 +49799,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_forInMacro_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(660u); +x_1 = lean_unsigned_to_nat(664u); x_2 = lean_unsigned_to_nat(89u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -49901,7 +49827,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_forInMacro_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(659u); +x_1 = lean_unsigned_to_nat(663u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -49913,7 +49839,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_forInMacro_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(659u); +x_1 = lean_unsigned_to_nat(663u); x_2 = lean_unsigned_to_nat(37u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -50292,7 +50218,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_forInMacro_x27_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(661u); +x_1 = lean_unsigned_to_nat(665u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -50304,7 +50230,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_forInMacro_x27_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(662u); +x_1 = lean_unsigned_to_nat(666u); x_2 = lean_unsigned_to_nat(90u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -50332,7 +50258,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_forInMacro_x27_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(661u); +x_1 = lean_unsigned_to_nat(665u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -50344,7 +50270,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_forInMacro_x27_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(661u); +x_1 = lean_unsigned_to_nat(665u); x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -50707,7 +50633,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_declName_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(665u); +x_1 = lean_unsigned_to_nat(669u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -50719,7 +50645,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_declName_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(665u); +x_1 = lean_unsigned_to_nat(669u); x_2 = lean_unsigned_to_nat(66u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -50747,7 +50673,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_declName_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(665u); +x_1 = lean_unsigned_to_nat(669u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -50759,7 +50685,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_declName_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(665u); +x_1 = lean_unsigned_to_nat(669u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -51038,7 +50964,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__8; +x_2 = l_Lean_Parser_Term_typeAscription___closed__9; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -51137,7 +51063,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_withDeclName_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(673u); +x_1 = lean_unsigned_to_nat(677u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -51149,7 +51075,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_withDeclName_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(674u); +x_1 = lean_unsigned_to_nat(678u); x_2 = lean_unsigned_to_nat(70u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -51177,7 +51103,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_withDeclName_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(673u); +x_1 = lean_unsigned_to_nat(677u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -51189,7 +51115,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_withDeclName_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(673u); +x_1 = lean_unsigned_to_nat(677u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -51600,7 +51526,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_typeOf_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(675u); +x_1 = lean_unsigned_to_nat(679u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -51612,7 +51538,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_typeOf_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(676u); +x_1 = lean_unsigned_to_nat(680u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -51640,7 +51566,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_typeOf_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(675u); +x_1 = lean_unsigned_to_nat(679u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -51652,7 +51578,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_typeOf_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(675u); +x_1 = lean_unsigned_to_nat(679u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -51946,7 +51872,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__8; +x_2 = l_Lean_Parser_Term_typeAscription___closed__9; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -52027,7 +51953,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_ensureTypeOf_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(677u); +x_1 = lean_unsigned_to_nat(681u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -52039,7 +51965,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_ensureTypeOf_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(678u); +x_1 = lean_unsigned_to_nat(682u); x_2 = lean_unsigned_to_nat(77u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -52067,7 +51993,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_ensureTypeOf_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(677u); +x_1 = lean_unsigned_to_nat(681u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -52079,7 +52005,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_ensureTypeOf_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(677u); +x_1 = lean_unsigned_to_nat(681u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -52492,7 +52418,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_ensureExpectedType_dec _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(679u); +x_1 = lean_unsigned_to_nat(683u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -52504,7 +52430,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_ensureExpectedType_dec _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(680u); +x_1 = lean_unsigned_to_nat(684u); x_2 = lean_unsigned_to_nat(69u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -52532,7 +52458,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_ensureExpectedType_dec _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(679u); +x_1 = lean_unsigned_to_nat(683u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -52544,7 +52470,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_ensureExpectedType_dec _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(679u); +x_1 = lean_unsigned_to_nat(683u); x_2 = lean_unsigned_to_nat(45u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -52911,7 +52837,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_noImplicitLambda_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(681u); +x_1 = lean_unsigned_to_nat(685u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -52923,7 +52849,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_noImplicitLambda_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(682u); +x_1 = lean_unsigned_to_nat(686u); x_2 = lean_unsigned_to_nat(46u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -52951,7 +52877,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_noImplicitLambda_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(681u); +x_1 = lean_unsigned_to_nat(685u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -52963,7 +52889,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_noImplicitLambda_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(681u); +x_1 = lean_unsigned_to_nat(685u); x_2 = lean_unsigned_to_nat(43u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -53257,7 +53183,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__8; +x_2 = l_Lean_Parser_Term_typeAscription___closed__9; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -53356,7 +53282,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_clear_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(688u); +x_1 = lean_unsigned_to_nat(692u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -53368,7 +53294,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_clear_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(689u); +x_1 = lean_unsigned_to_nat(693u); x_2 = lean_unsigned_to_nat(77u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -53396,7 +53322,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_clear_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(688u); +x_1 = lean_unsigned_to_nat(692u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -53408,7 +53334,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_clear_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(688u); +x_1 = lean_unsigned_to_nat(692u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -53774,7 +53700,7 @@ static lean_object* _init_l_Lean_Parser_Term_letMVar___closed__6() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_sepByIndentSemicolon___closed__1; -x_2 = l_Lean_Parser_Termination_terminationBy___closed__15; +x_2 = l_Lean_Parser_Term_typeAscription___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -53783,7 +53709,7 @@ static lean_object* _init_l_Lean_Parser_Term_letMVar___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Termination_terminationBy___closed__15; +x_1 = l_Lean_Parser_Term_typeAscription___closed__6; x_2 = l_Lean_Parser_Term_letMVar___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -53885,7 +53811,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_letMVar_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(691u); +x_1 = lean_unsigned_to_nat(695u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -53897,7 +53823,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_letMVar_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(692u); +x_1 = lean_unsigned_to_nat(696u); x_2 = lean_unsigned_to_nat(76u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -53925,7 +53851,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_letMVar_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(691u); +x_1 = lean_unsigned_to_nat(695u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -53937,7 +53863,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_letMVar_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(691u); +x_1 = lean_unsigned_to_nat(695u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -54351,7 +54277,7 @@ static lean_object* _init_l_Lean_Parser_Term_waitIfTypeMVar___closed__6() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_sepByIndentSemicolon___closed__1; -x_2 = l_Lean_Parser_Termination_terminationBy___closed__15; +x_2 = l_Lean_Parser_Term_typeAscription___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -54442,7 +54368,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_waitIfTypeMVar_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(693u); +x_1 = lean_unsigned_to_nat(697u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -54454,7 +54380,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_waitIfTypeMVar_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(694u); +x_1 = lean_unsigned_to_nat(698u); x_2 = lean_unsigned_to_nat(61u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -54482,7 +54408,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_waitIfTypeMVar_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(693u); +x_1 = lean_unsigned_to_nat(697u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -54494,7 +54420,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_waitIfTypeMVar_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(693u); +x_1 = lean_unsigned_to_nat(697u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -54909,7 +54835,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_waitIfTypeContainsMVar _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(695u); +x_1 = lean_unsigned_to_nat(699u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -54921,7 +54847,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_waitIfTypeContainsMVar _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(696u); +x_1 = lean_unsigned_to_nat(700u); x_2 = lean_unsigned_to_nat(70u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -54949,7 +54875,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_waitIfTypeContainsMVar _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(695u); +x_1 = lean_unsigned_to_nat(699u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -54961,7 +54887,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_waitIfTypeContainsMVar _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(695u); +x_1 = lean_unsigned_to_nat(699u); x_2 = lean_unsigned_to_nat(49u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -55316,7 +55242,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_waitIfContainsMVar_dec _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(697u); +x_1 = lean_unsigned_to_nat(701u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -55328,7 +55254,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_waitIfContainsMVar_dec _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(698u); +x_1 = lean_unsigned_to_nat(702u); x_2 = lean_unsigned_to_nat(65u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -55356,7 +55282,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_waitIfContainsMVar_dec _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(697u); +x_1 = lean_unsigned_to_nat(701u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -55368,7 +55294,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_waitIfContainsMVar_dec _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(697u); +x_1 = lean_unsigned_to_nat(701u); x_2 = lean_unsigned_to_nat(45u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -55741,7 +55667,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_defaultOrOfNonempty_de _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(700u); +x_1 = lean_unsigned_to_nat(704u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -55753,7 +55679,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_defaultOrOfNonempty_de _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(701u); +x_1 = lean_unsigned_to_nat(705u); x_2 = lean_unsigned_to_nat(48u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -55781,7 +55707,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_defaultOrOfNonempty_de _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(700u); +x_1 = lean_unsigned_to_nat(704u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -55793,7 +55719,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_defaultOrOfNonempty_de _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(700u); +x_1 = lean_unsigned_to_nat(704u); x_2 = lean_unsigned_to_nat(46u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -56127,7 +56053,7 @@ static lean_object* _init_l_Lean_Parser_Term_noErrorIfUnused___closed__6() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_noErrorIfUnused___closed__5; -x_2 = l_Lean_Parser_Termination_terminationBy___closed__15; +x_2 = l_Lean_Parser_Term_typeAscription___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -56206,7 +56132,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_noErrorIfUnused_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(707u); +x_1 = lean_unsigned_to_nat(711u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -56218,7 +56144,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_noErrorIfUnused_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(708u); +x_1 = lean_unsigned_to_nat(712u); x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -56246,7 +56172,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_noErrorIfUnused_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(707u); +x_1 = lean_unsigned_to_nat(711u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -56258,7 +56184,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_noErrorIfUnused_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(707u); +x_1 = lean_unsigned_to_nat(711u); x_2 = lean_unsigned_to_nat(42u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -56879,7 +56805,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_app_declRange___closed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(721u); +x_1 = lean_unsigned_to_nat(725u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -56891,7 +56817,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_app_declRange___closed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(721u); +x_1 = lean_unsigned_to_nat(725u); x_2 = lean_unsigned_to_nat(81u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -56919,7 +56845,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_app_declRange___closed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(721u); +x_1 = lean_unsigned_to_nat(725u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -56931,7 +56857,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_app_declRange___closed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(721u); +x_1 = lean_unsigned_to_nat(725u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -57772,7 +57698,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_proj_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(735u); +x_1 = lean_unsigned_to_nat(739u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -57784,7 +57710,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_proj_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(736u); +x_1 = lean_unsigned_to_nat(740u); x_2 = lean_unsigned_to_nat(70u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -57812,7 +57738,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_proj_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(735u); +x_1 = lean_unsigned_to_nat(739u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -57824,7 +57750,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_proj_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(735u); +x_1 = lean_unsigned_to_nat(739u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -58139,7 +58065,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_completion_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(737u); +x_1 = lean_unsigned_to_nat(741u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -58151,7 +58077,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_completion_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(738u); +x_1 = lean_unsigned_to_nat(742u); x_2 = lean_unsigned_to_nat(24u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -58179,7 +58105,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_completion_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(737u); +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); @@ -58191,7 +58117,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_completion_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(737u); +x_1 = lean_unsigned_to_nat(741u); x_2 = lean_unsigned_to_nat(37u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -58446,7 +58372,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_arrow_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(739u); +x_1 = lean_unsigned_to_nat(743u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -58458,7 +58384,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_arrow_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(740u); +x_1 = lean_unsigned_to_nat(744u); x_2 = lean_unsigned_to_nat(61u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -58486,7 +58412,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_arrow_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(739u); +x_1 = lean_unsigned_to_nat(743u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -58498,7 +58424,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_arrow_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(739u); +x_1 = lean_unsigned_to_nat(743u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -58916,7 +58842,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_explicitUniv_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(747u); +x_1 = lean_unsigned_to_nat(751u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -58928,7 +58854,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_explicitUniv_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(750u); +x_1 = lean_unsigned_to_nat(754u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -58956,7 +58882,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_explicitUniv_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(747u); +x_1 = lean_unsigned_to_nat(751u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -58968,7 +58894,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_explicitUniv_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(747u); +x_1 = lean_unsigned_to_nat(751u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -59440,7 +59366,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_namedPattern_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(753u); +x_1 = lean_unsigned_to_nat(757u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -59452,7 +59378,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_namedPattern_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(756u); +x_1 = lean_unsigned_to_nat(760u); x_2 = lean_unsigned_to_nat(56u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -59480,7 +59406,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_namedPattern_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(753u); +x_1 = lean_unsigned_to_nat(757u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -59492,7 +59418,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_namedPattern_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(753u); +x_1 = lean_unsigned_to_nat(757u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -59919,7 +59845,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_pipeProj_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(762u); +x_1 = lean_unsigned_to_nat(766u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -59931,7 +59857,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_pipeProj_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(763u); +x_1 = lean_unsigned_to_nat(767u); x_2 = lean_unsigned_to_nat(71u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -59959,7 +59885,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_pipeProj_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(762u); +x_1 = lean_unsigned_to_nat(766u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -59971,7 +59897,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_pipeProj_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(762u); +x_1 = lean_unsigned_to_nat(766u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -60276,7 +60202,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_pipeCompletion_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(764u); +x_1 = lean_unsigned_to_nat(768u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -60288,7 +60214,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_pipeCompletion_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(765u); +x_1 = lean_unsigned_to_nat(769u); x_2 = lean_unsigned_to_nat(8u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -60316,7 +60242,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_pipeCompletion_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(764u); +x_1 = lean_unsigned_to_nat(768u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -60328,7 +60254,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_pipeCompletion_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(764u); +x_1 = lean_unsigned_to_nat(768u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -60587,7 +60513,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_subst_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(775u); +x_1 = lean_unsigned_to_nat(779u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -60599,7 +60525,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_subst_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(776u); +x_1 = lean_unsigned_to_nat(780u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -60627,7 +60553,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_subst_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(775u); +x_1 = lean_unsigned_to_nat(779u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -60639,7 +60565,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_subst_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(775u); +x_1 = lean_unsigned_to_nat(779u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -60941,7 +60867,7 @@ static lean_object* _init_l_Lean_Parser_Term_panic___closed__6() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_panic___closed__5; -x_2 = l_Lean_Parser_Termination_terminationBy___closed__15; +x_2 = l_Lean_Parser_Term_typeAscription___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -61020,7 +60946,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_panic_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(789u); +x_1 = lean_unsigned_to_nat(793u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -61032,7 +60958,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_panic_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(790u); +x_1 = lean_unsigned_to_nat(794u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -61060,7 +60986,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_panic_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(789u); +x_1 = lean_unsigned_to_nat(793u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -61072,7 +60998,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_panic_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(789u); +x_1 = lean_unsigned_to_nat(793u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -61435,7 +61361,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_unreachable_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(792u); +x_1 = lean_unsigned_to_nat(796u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -61447,7 +61373,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_unreachable_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(793u); +x_1 = lean_unsigned_to_nat(797u); x_2 = lean_unsigned_to_nat(16u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -61475,7 +61401,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_unreachable_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(792u); +x_1 = lean_unsigned_to_nat(796u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -61487,7 +61413,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_unreachable_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(792u); +x_1 = lean_unsigned_to_nat(796u); x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -61756,7 +61682,7 @@ static lean_object* _init_l_Lean_Parser_Term_dbgTrace___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Termination_terminationBy___closed__15; +x_1 = l_Lean_Parser_Term_typeAscription___closed__6; x_2 = l_Lean_Parser_interpolatedStr(x_1); return x_2; } @@ -61766,7 +61692,7 @@ static lean_object* _init_l_Lean_Parser_Term_dbgTrace___closed__7() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_dbgTrace___closed__6; -x_2 = l_Lean_Parser_Termination_terminationBy___closed__15; +x_2 = l_Lean_Parser_Term_typeAscription___closed__6; x_3 = l_Lean_Parser_orelse(x_1, x_2); return x_3; } @@ -61874,7 +61800,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_dbgTrace_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(798u); +x_1 = lean_unsigned_to_nat(802u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -61886,7 +61812,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_dbgTrace_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(800u); +x_1 = lean_unsigned_to_nat(804u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -61914,7 +61840,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_dbgTrace_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(798u); +x_1 = lean_unsigned_to_nat(802u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -61926,7 +61852,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_dbgTrace_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(798u); +x_1 = lean_unsigned_to_nat(802u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -62308,7 +62234,7 @@ static lean_object* _init_l_Lean_Parser_Term_assert___closed__6() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_assert___closed__5; -x_2 = l_Lean_Parser_Termination_terminationBy___closed__15; +x_2 = l_Lean_Parser_Term_typeAscription___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -62406,7 +62332,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_assert_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(802u); +x_1 = lean_unsigned_to_nat(806u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -62418,7 +62344,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_assert_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(803u); +x_1 = lean_unsigned_to_nat(807u); x_2 = lean_unsigned_to_nat(68u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -62446,7 +62372,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_assert_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(802u); +x_1 = lean_unsigned_to_nat(806u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -62458,7 +62384,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_assert_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(802u); +x_1 = lean_unsigned_to_nat(806u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -63010,7 +62936,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_stateRefT_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(810u); +x_1 = lean_unsigned_to_nat(814u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -63022,7 +62948,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_stateRefT_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(811u); +x_1 = lean_unsigned_to_nat(815u); x_2 = lean_unsigned_to_nat(53u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -63050,7 +62976,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_stateRefT_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(810u); +x_1 = lean_unsigned_to_nat(814u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -63062,7 +62988,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_stateRefT_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(810u); +x_1 = lean_unsigned_to_nat(814u); x_2 = lean_unsigned_to_nat(36u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -63696,7 +63622,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; } @@ -63766,7 +63692,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; @@ -63797,7 +63723,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_dynamicQuot_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(813u); +x_1 = lean_unsigned_to_nat(817u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -63809,7 +63735,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_dynamicQuot_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(814u); +x_1 = lean_unsigned_to_nat(818u); x_2 = lean_unsigned_to_nat(64u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -63837,7 +63763,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_dynamicQuot_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(813u); +x_1 = lean_unsigned_to_nat(817u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -63849,7 +63775,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_dynamicQuot_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(813u); +x_1 = lean_unsigned_to_nat(817u); x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -64313,7 +64239,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_dotIdent_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(816u); +x_1 = lean_unsigned_to_nat(820u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -64325,7 +64251,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_dotIdent_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(817u); +x_1 = lean_unsigned_to_nat(821u); x_2 = lean_unsigned_to_nat(36u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -64353,7 +64279,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_dotIdent_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(816u); +x_1 = lean_unsigned_to_nat(820u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -64365,7 +64291,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_dotIdent_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(816u); +x_1 = lean_unsigned_to_nat(820u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -64663,7 +64589,7 @@ static lean_object* _init_l_Lean_Parser_Term_showTermElabImpl___closed__6() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_showTermElabImpl___closed__5; -x_2 = l_Lean_Parser_Termination_terminationBy___closed__15; +x_2 = l_Lean_Parser_Term_typeAscription___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -64742,7 +64668,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_showTermElabImpl_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(822u); +x_1 = lean_unsigned_to_nat(826u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -64754,7 +64680,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_showTermElabImpl_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(823u); +x_1 = lean_unsigned_to_nat(827u); x_2 = lean_unsigned_to_nat(57u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -64782,7 +64708,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_showTermElabImpl_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(822u); +x_1 = lean_unsigned_to_nat(826u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -64794,7 +64720,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_showTermElabImpl_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(822u); +x_1 = lean_unsigned_to_nat(826u); x_2 = lean_unsigned_to_nat(43u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -65410,7 +65336,7 @@ static lean_object* _init_l_Lean_Parser_Term_matchExpr___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Termination_terminationBy___closed__15; +x_1 = l_Lean_Parser_Term_typeAscription___closed__6; x_2 = l_Lean_Parser_Term_matchExprAlts(x_1); return x_2; } @@ -65429,7 +65355,7 @@ static lean_object* _init_l_Lean_Parser_Term_matchExpr___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Termination_terminationBy___closed__15; +x_1 = l_Lean_Parser_Term_typeAscription___closed__6; x_2 = l_Lean_Parser_Term_matchExpr___closed__7; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -65501,7 +65427,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_matchExpr_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(836u); +x_1 = lean_unsigned_to_nat(840u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -65513,7 +65439,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_matchExpr_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(837u); +x_1 = lean_unsigned_to_nat(841u); x_2 = lean_unsigned_to_nat(79u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -65541,7 +65467,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_matchExpr_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(836u); +x_1 = lean_unsigned_to_nat(840u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -65553,7 +65479,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_matchExpr_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(836u); +x_1 = lean_unsigned_to_nat(840u); x_2 = lean_unsigned_to_nat(36u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -65762,7 +65688,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_matchExprAlt_formatter(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; -x_7 = l_Lean_Parser_Termination_terminationBy_formatter___closed__6; +x_7 = l_Lean_Parser_Termination_terminationBy_formatter___closed__4; x_8 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_8, 0, x_7); lean_closure_set(x_8, 1, x_1); @@ -65809,7 +65735,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_matchExprElseAlt_formatter(lean_obje _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; -x_7 = l_Lean_Parser_Termination_terminationBy_formatter___closed__6; +x_7 = l_Lean_Parser_Termination_terminationBy_formatter___closed__4; x_8 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_8, 0, x_7); lean_closure_set(x_8, 1, x_1); @@ -66233,7 +66159,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_matchExprAlt_parenthesizer(lean_obje _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; -x_7 = l_Lean_Parser_Termination_terminationBy_parenthesizer___closed__6; +x_7 = l_Lean_Parser_Termination_terminationBy_parenthesizer___closed__4; x_8 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_8, 0, x_7); lean_closure_set(x_8, 1, x_1); @@ -66280,7 +66206,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_matchExprElseAlt_parenthesizer(lean_ _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; -x_7 = l_Lean_Parser_Termination_terminationBy_parenthesizer___closed__6; +x_7 = l_Lean_Parser_Termination_terminationBy_parenthesizer___closed__4; x_8 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_8, 0, x_7); lean_closure_set(x_8, 1, x_1); @@ -66595,7 +66521,7 @@ static lean_object* _init_l_Lean_Parser_Term_letExpr___closed__6() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_matchAlt___closed__7; -x_2 = l_Lean_Parser_Termination_terminationBy___closed__15; +x_2 = l_Lean_Parser_Term_typeAscription___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -66614,7 +66540,7 @@ static lean_object* _init_l_Lean_Parser_Term_letExpr___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Termination_terminationBy___closed__15; +x_1 = l_Lean_Parser_Term_typeAscription___closed__6; x_2 = l_Lean_Parser_Term_letExpr___closed__7; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -66725,7 +66651,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_letExpr_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(839u); +x_1 = lean_unsigned_to_nat(843u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -66737,7 +66663,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_letExpr_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(840u); +x_1 = lean_unsigned_to_nat(844u); x_2 = lean_unsigned_to_nat(132u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -66765,7 +66691,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_letExpr_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(839u); +x_1 = lean_unsigned_to_nat(843u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -66777,7 +66703,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_letExpr_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(839u); +x_1 = lean_unsigned_to_nat(843u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -67243,7 +67169,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; @@ -67254,7 +67180,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; } @@ -67325,7 +67251,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_quot_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(844u); +x_1 = lean_unsigned_to_nat(848u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -67337,7 +67263,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_quot_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(845u); +x_1 = lean_unsigned_to_nat(849u); x_2 = lean_unsigned_to_nat(68u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -67365,7 +67291,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_quot_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(844u); +x_1 = lean_unsigned_to_nat(848u); x_2 = lean_unsigned_to_nat(37u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -67377,7 +67303,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_quot_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(844u); +x_1 = lean_unsigned_to_nat(848u); x_2 = lean_unsigned_to_nat(48u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -67726,7 +67652,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; @@ -67737,7 +67663,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; } @@ -67808,7 +67734,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_quotSeq_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(846u); +x_1 = lean_unsigned_to_nat(850u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -67820,7 +67746,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_quotSeq_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(847u); +x_1 = lean_unsigned_to_nat(851u); x_2 = lean_unsigned_to_nat(67u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -67848,7 +67774,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_quotSeq_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(846u); +x_1 = lean_unsigned_to_nat(850u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -67860,7 +67786,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_quotSeq_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(846u); +x_1 = lean_unsigned_to_nat(850u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -68284,7 +68210,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_5127____closed__1() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -68294,7 +68220,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_5127____closed__2() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__2() { _start: { lean_object* x_1; lean_object* x_2; @@ -68304,7 +68230,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_5127____closed__3() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__3() { _start: { lean_object* x_1; lean_object* x_2; @@ -68314,7 +68240,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_5127____closed__4() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__4() { _start: { lean_object* x_1; lean_object* x_2; @@ -68324,7 +68250,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_5127____closed__5() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__5() { _start: { lean_object* x_1; lean_object* x_2; @@ -68334,7 +68260,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_5127____closed__6() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -68344,7 +68270,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_5127____closed__7() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__7() { _start: { lean_object* x_1; lean_object* x_2; @@ -68354,7 +68280,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_5127____closed__8() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__8() { _start: { lean_object* x_1; lean_object* x_2; @@ -68364,7 +68290,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_5127____closed__9() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__9() { _start: { lean_object* x_1; lean_object* x_2; @@ -68374,7 +68300,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_5127____closed__10() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__10() { _start: { lean_object* x_1; lean_object* x_2; @@ -68384,7 +68310,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_5127____closed__11() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -68394,7 +68320,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_5127____closed__12() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__12() { _start: { lean_object* x_1; lean_object* x_2; @@ -68404,7 +68330,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_5127____closed__13() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__13() { _start: { lean_object* x_1; lean_object* x_2; @@ -68414,7 +68340,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_5127____closed__14() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__14() { _start: { lean_object* x_1; lean_object* x_2; @@ -68424,7 +68350,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_5127____closed__15() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__15() { _start: { lean_object* x_1; lean_object* x_2; @@ -68434,7 +68360,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_5127____closed__16() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -68444,7 +68370,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_5127____closed__17() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__17() { _start: { lean_object* x_1; lean_object* x_2; @@ -68454,7 +68380,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_5127____closed__18() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__18() { _start: { lean_object* x_1; lean_object* x_2; @@ -68464,7 +68390,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_5127____closed__19() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__19() { _start: { lean_object* x_1; lean_object* x_2; @@ -68474,7 +68400,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_5127____closed__20() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__20() { _start: { lean_object* x_1; lean_object* x_2; @@ -68484,7 +68410,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_5127____closed__21() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -68494,7 +68420,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_5127____closed__22() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__22() { _start: { lean_object* x_1; lean_object* x_2; @@ -68504,7 +68430,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_5127____closed__23() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__23() { _start: { lean_object* x_1; lean_object* x_2; @@ -68514,7 +68440,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_5127____closed__24() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__24() { _start: { lean_object* x_1; lean_object* x_2; @@ -68524,7 +68450,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_5127____closed__25() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__25() { _start: { lean_object* x_1; lean_object* x_2; @@ -68534,7 +68460,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_5127____closed__26() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__26() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -68544,7 +68470,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_5127____closed__27() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__27() { _start: { lean_object* x_1; lean_object* x_2; @@ -68554,7 +68480,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_5127____closed__28() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__28() { _start: { lean_object* x_1; lean_object* x_2; @@ -68564,7 +68490,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_5127____closed__29() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__29() { _start: { lean_object* x_1; lean_object* x_2; @@ -68574,7 +68500,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_5127____closed__30() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__30() { _start: { lean_object* x_1; lean_object* x_2; @@ -68584,7 +68510,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_5127____closed__31() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__31() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -68594,7 +68520,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_5127____closed__32() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__32() { _start: { lean_object* x_1; lean_object* x_2; @@ -68604,7 +68530,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_5127____closed__33() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__33() { _start: { lean_object* x_1; lean_object* x_2; @@ -68614,7 +68540,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_5127____closed__34() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__34() { _start: { lean_object* x_1; lean_object* x_2; @@ -68624,7 +68550,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_5127____closed__35() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__35() { _start: { lean_object* x_1; lean_object* x_2; @@ -68634,7 +68560,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_5127____closed__36() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__36() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -68644,7 +68570,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_5127____closed__37() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__37() { _start: { lean_object* x_1; lean_object* x_2; @@ -68654,7 +68580,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_5127____closed__38() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__38() { _start: { lean_object* x_1; lean_object* x_2; @@ -68664,7 +68590,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_5127____closed__39() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__39() { _start: { lean_object* x_1; lean_object* x_2; @@ -68674,7 +68600,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_5127____closed__40() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__40() { _start: { lean_object* x_1; lean_object* x_2; @@ -68684,7 +68610,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_5127____closed__41() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__41() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -68694,7 +68620,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_5127____closed__42() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__42() { _start: { lean_object* x_1; lean_object* x_2; @@ -68704,7 +68630,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_5127____closed__43() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__43() { _start: { lean_object* x_1; lean_object* x_2; @@ -68714,7 +68640,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_5127____closed__44() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__44() { _start: { lean_object* x_1; lean_object* x_2; @@ -68724,7 +68650,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_5127____closed__45() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__45() { _start: { lean_object* x_1; lean_object* x_2; @@ -68734,7 +68660,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_5127____closed__46() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__46() { _start: { lean_object* x_1; @@ -68742,29 +68668,29 @@ x_1 = lean_mk_string_from_bytes("optSemicolon", 12); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__47() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____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_Term___hyg_5127____closed__46; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____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_Term___hyg_5127____closed__48() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__48() { _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_Command_docComment___closed__1; x_2 = l_Lean_Parser_Command_docComment___closed__2; x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_66____closed__17; -x_4 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__46; +x_4 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____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_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__49() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__49() { _start: { lean_object* x_1; @@ -68772,27 +68698,27 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_optSemicolon), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__50() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__50() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__49; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____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_Term___hyg_5127____closed__51() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__51() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__48; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____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_Term___hyg_5127____closed__52() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__52() { _start: { lean_object* x_1; @@ -68800,17 +68726,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_optSemicolon_formatter), 6, return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__53() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__53() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__52; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__52; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__54() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__54() { _start: { lean_object* x_1; @@ -68818,24 +68744,24 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_optSemicolon_parenthesizer), return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__55() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__55() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__54; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__54; 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_Parser_initFn____x40_Lean_Parser_Term___hyg_5127_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126_(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_5127____closed__1; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__1; x_3 = l_Lean_Parser_Term_letDecl___closed__2; -x_4 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__2; -x_5 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__3; +x_4 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__2; +x_5 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__3; x_6 = l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Term___hyg_154____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) @@ -68845,7 +68771,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_154____closed__12; -x_10 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__4; +x_10 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__4; x_11 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_2, x_10, x_8); if (lean_obj_tag(x_11) == 0) { @@ -68854,7 +68780,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_154____closed__15; -x_14 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__5; +x_14 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__5; x_15 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_2, x_14, x_12); if (lean_obj_tag(x_15) == 0) { @@ -68862,10 +68788,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_5127____closed__6; +x_17 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__6; x_18 = l_Lean_Parser_Term_haveDecl___closed__2; -x_19 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__7; -x_20 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__8; +x_19 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__7; +x_20 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____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) { @@ -68873,7 +68799,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_5127____closed__9; +x_23 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__9; x_24 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_17, x_23, x_22); if (lean_obj_tag(x_24) == 0) { @@ -68881,7 +68807,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_5127____closed__10; +x_26 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__10; x_27 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_17, x_26, x_25); if (lean_obj_tag(x_27) == 0) { @@ -68889,10 +68815,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_5127____closed__11; +x_29 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__11; x_30 = l_Lean_Parser_Term_sufficesDecl___closed__2; -x_31 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__12; -x_32 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__13; +x_31 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__12; +x_32 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____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) { @@ -68900,7 +68826,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_5127____closed__14; +x_35 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__14; x_36 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_29, x_35, x_34); if (lean_obj_tag(x_36) == 0) { @@ -68908,7 +68834,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_5127____closed__15; +x_38 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__15; x_39 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_29, x_38, x_37); if (lean_obj_tag(x_39) == 0) { @@ -68916,10 +68842,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_5127____closed__16; +x_41 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__16; x_42 = l_Lean_Parser_Term_letRecDecls___closed__2; -x_43 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__17; -x_44 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__18; +x_43 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__17; +x_44 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____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) { @@ -68927,7 +68853,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_5127____closed__19; +x_47 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__19; x_48 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_41, x_47, x_46); if (lean_obj_tag(x_48) == 0) { @@ -68935,7 +68861,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_5127____closed__20; +x_50 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__20; x_51 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_41, x_50, x_49); if (lean_obj_tag(x_51) == 0) { @@ -68943,10 +68869,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_5127____closed__21; +x_53 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__21; x_54 = l_Lean_Parser_Term_hole___closed__2; -x_55 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__22; -x_56 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__23; +x_55 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__22; +x_56 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____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) { @@ -68954,7 +68880,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_5127____closed__24; +x_59 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__24; x_60 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_53, x_59, x_58); if (lean_obj_tag(x_60) == 0) { @@ -68962,7 +68888,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_5127____closed__25; +x_62 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__25; x_63 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_53, x_62, x_61); if (lean_obj_tag(x_63) == 0) { @@ -68970,10 +68896,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_5127____closed__26; +x_65 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__26; x_66 = l_Lean_Parser_Term_syntheticHole___closed__2; -x_67 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__27; -x_68 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__28; +x_67 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__27; +x_68 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____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) { @@ -68981,7 +68907,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_5127____closed__29; +x_71 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__29; x_72 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_65, x_71, x_70); if (lean_obj_tag(x_72) == 0) { @@ -68989,7 +68915,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_5127____closed__30; +x_74 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__30; x_75 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_65, x_74, x_73); if (lean_obj_tag(x_75) == 0) { @@ -68997,10 +68923,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_5127____closed__31; +x_77 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__31; x_78 = l_Lean_Parser_Term_matchDiscr___closed__2; -x_79 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__32; -x_80 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__33; +x_79 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__32; +x_80 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____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) { @@ -69008,7 +68934,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_5127____closed__34; +x_83 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__34; x_84 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_77, x_83, x_82); if (lean_obj_tag(x_84) == 0) { @@ -69016,7 +68942,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_5127____closed__35; +x_86 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__35; x_87 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_77, x_86, x_85); if (lean_obj_tag(x_87) == 0) { @@ -69024,10 +68950,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_5127____closed__36; +x_89 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__36; x_90 = l_Lean_Parser_Term_bracketedBinder___closed__2; -x_91 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__37; -x_92 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__38; +x_91 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__37; +x_92 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____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) { @@ -69035,7 +68961,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_5127____closed__39; +x_95 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__39; x_96 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_89, x_95, x_94); if (lean_obj_tag(x_96) == 0) { @@ -69043,7 +68969,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_5127____closed__40; +x_98 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__40; x_99 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_89, x_98, x_97); if (lean_obj_tag(x_99) == 0) { @@ -69051,10 +68977,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_5127____closed__41; +x_101 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__41; x_102 = l_Lean_Parser_Term_attrKind___closed__2; -x_103 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__42; -x_104 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__43; +x_103 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__42; +x_104 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____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) { @@ -69062,7 +68988,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_5127____closed__44; +x_107 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__44; x_108 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_101, x_107, x_106); if (lean_obj_tag(x_108) == 0) { @@ -69070,7 +68996,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_5127____closed__45; +x_110 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__45; x_111 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_101, x_110, x_109); if (lean_obj_tag(x_111) == 0) { @@ -69078,10 +69004,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_Term___hyg_5127____closed__47; -x_114 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__48; -x_115 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__50; -x_116 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__51; +x_113 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__47; +x_114 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__48; +x_115 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__50; +x_116 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__51; x_117 = l_Lean_Parser_registerAlias(x_113, x_114, x_115, x_116, x_6, x_112); if (lean_obj_tag(x_117) == 0) { @@ -69089,7 +69015,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_Term___hyg_5127____closed__53; +x_119 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__53; x_120 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_113, x_119, x_118); if (lean_obj_tag(x_120) == 0) { @@ -69097,7 +69023,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_Term___hyg_5127____closed__55; +x_122 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__55; x_123 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_113, x_122, x_121); return x_123; } @@ -70064,112 +69990,6 @@ l_Lean_Parser_semicolonOrLinebreak___closed__3 = _init_l_Lean_Parser_semicolonOr lean_mark_persistent(l_Lean_Parser_semicolonOrLinebreak___closed__3); l_Lean_Parser_semicolonOrLinebreak = _init_l_Lean_Parser_semicolonOrLinebreak(); lean_mark_persistent(l_Lean_Parser_semicolonOrLinebreak); -l_Lean_Parser_Termination_terminationBy___closed__1 = _init_l_Lean_Parser_Termination_terminationBy___closed__1(); -lean_mark_persistent(l_Lean_Parser_Termination_terminationBy___closed__1); -l_Lean_Parser_Termination_terminationBy___closed__2 = _init_l_Lean_Parser_Termination_terminationBy___closed__2(); -lean_mark_persistent(l_Lean_Parser_Termination_terminationBy___closed__2); -l_Lean_Parser_Termination_terminationBy___closed__3 = _init_l_Lean_Parser_Termination_terminationBy___closed__3(); -lean_mark_persistent(l_Lean_Parser_Termination_terminationBy___closed__3); -l_Lean_Parser_Termination_terminationBy___closed__4 = _init_l_Lean_Parser_Termination_terminationBy___closed__4(); -lean_mark_persistent(l_Lean_Parser_Termination_terminationBy___closed__4); -l_Lean_Parser_Termination_terminationBy___closed__5 = _init_l_Lean_Parser_Termination_terminationBy___closed__5(); -lean_mark_persistent(l_Lean_Parser_Termination_terminationBy___closed__5); -l_Lean_Parser_Termination_terminationBy___closed__6 = _init_l_Lean_Parser_Termination_terminationBy___closed__6(); -lean_mark_persistent(l_Lean_Parser_Termination_terminationBy___closed__6); -l_Lean_Parser_Termination_terminationBy___closed__7 = _init_l_Lean_Parser_Termination_terminationBy___closed__7(); -lean_mark_persistent(l_Lean_Parser_Termination_terminationBy___closed__7); -l_Lean_Parser_Termination_terminationBy___closed__8 = _init_l_Lean_Parser_Termination_terminationBy___closed__8(); -lean_mark_persistent(l_Lean_Parser_Termination_terminationBy___closed__8); -l_Lean_Parser_Termination_terminationBy___closed__9 = _init_l_Lean_Parser_Termination_terminationBy___closed__9(); -lean_mark_persistent(l_Lean_Parser_Termination_terminationBy___closed__9); -l_Lean_Parser_Termination_terminationBy___closed__10 = _init_l_Lean_Parser_Termination_terminationBy___closed__10(); -lean_mark_persistent(l_Lean_Parser_Termination_terminationBy___closed__10); -l_Lean_Parser_Termination_terminationBy___closed__11 = _init_l_Lean_Parser_Termination_terminationBy___closed__11(); -lean_mark_persistent(l_Lean_Parser_Termination_terminationBy___closed__11); -l_Lean_Parser_Termination_terminationBy___closed__12 = _init_l_Lean_Parser_Termination_terminationBy___closed__12(); -lean_mark_persistent(l_Lean_Parser_Termination_terminationBy___closed__12); -l_Lean_Parser_Termination_terminationBy___closed__13 = _init_l_Lean_Parser_Termination_terminationBy___closed__13(); -lean_mark_persistent(l_Lean_Parser_Termination_terminationBy___closed__13); -l_Lean_Parser_Termination_terminationBy___closed__14 = _init_l_Lean_Parser_Termination_terminationBy___closed__14(); -lean_mark_persistent(l_Lean_Parser_Termination_terminationBy___closed__14); -l_Lean_Parser_Termination_terminationBy___closed__15 = _init_l_Lean_Parser_Termination_terminationBy___closed__15(); -lean_mark_persistent(l_Lean_Parser_Termination_terminationBy___closed__15); -l_Lean_Parser_Termination_terminationBy___closed__16 = _init_l_Lean_Parser_Termination_terminationBy___closed__16(); -lean_mark_persistent(l_Lean_Parser_Termination_terminationBy___closed__16); -l_Lean_Parser_Termination_terminationBy___closed__17 = _init_l_Lean_Parser_Termination_terminationBy___closed__17(); -lean_mark_persistent(l_Lean_Parser_Termination_terminationBy___closed__17); -l_Lean_Parser_Termination_terminationBy___closed__18 = _init_l_Lean_Parser_Termination_terminationBy___closed__18(); -lean_mark_persistent(l_Lean_Parser_Termination_terminationBy___closed__18); -l_Lean_Parser_Termination_terminationBy___closed__19 = _init_l_Lean_Parser_Termination_terminationBy___closed__19(); -lean_mark_persistent(l_Lean_Parser_Termination_terminationBy___closed__19); -l_Lean_Parser_Termination_terminationBy___closed__20 = _init_l_Lean_Parser_Termination_terminationBy___closed__20(); -lean_mark_persistent(l_Lean_Parser_Termination_terminationBy___closed__20); -l_Lean_Parser_Termination_terminationBy = _init_l_Lean_Parser_Termination_terminationBy(); -lean_mark_persistent(l_Lean_Parser_Termination_terminationBy); -l_Lean_Parser_Termination_terminationBy_x3f___closed__1 = _init_l_Lean_Parser_Termination_terminationBy_x3f___closed__1(); -lean_mark_persistent(l_Lean_Parser_Termination_terminationBy_x3f___closed__1); -l_Lean_Parser_Termination_terminationBy_x3f___closed__2 = _init_l_Lean_Parser_Termination_terminationBy_x3f___closed__2(); -lean_mark_persistent(l_Lean_Parser_Termination_terminationBy_x3f___closed__2); -l_Lean_Parser_Termination_terminationBy_x3f___closed__3 = _init_l_Lean_Parser_Termination_terminationBy_x3f___closed__3(); -lean_mark_persistent(l_Lean_Parser_Termination_terminationBy_x3f___closed__3); -l_Lean_Parser_Termination_terminationBy_x3f___closed__4 = _init_l_Lean_Parser_Termination_terminationBy_x3f___closed__4(); -lean_mark_persistent(l_Lean_Parser_Termination_terminationBy_x3f___closed__4); -l_Lean_Parser_Termination_terminationBy_x3f___closed__5 = _init_l_Lean_Parser_Termination_terminationBy_x3f___closed__5(); -lean_mark_persistent(l_Lean_Parser_Termination_terminationBy_x3f___closed__5); -l_Lean_Parser_Termination_terminationBy_x3f___closed__6 = _init_l_Lean_Parser_Termination_terminationBy_x3f___closed__6(); -lean_mark_persistent(l_Lean_Parser_Termination_terminationBy_x3f___closed__6); -l_Lean_Parser_Termination_terminationBy_x3f___closed__7 = _init_l_Lean_Parser_Termination_terminationBy_x3f___closed__7(); -lean_mark_persistent(l_Lean_Parser_Termination_terminationBy_x3f___closed__7); -l_Lean_Parser_Termination_terminationBy_x3f___closed__8 = _init_l_Lean_Parser_Termination_terminationBy_x3f___closed__8(); -lean_mark_persistent(l_Lean_Parser_Termination_terminationBy_x3f___closed__8); -l_Lean_Parser_Termination_terminationBy_x3f = _init_l_Lean_Parser_Termination_terminationBy_x3f(); -lean_mark_persistent(l_Lean_Parser_Termination_terminationBy_x3f); -l_Lean_Parser_Termination_decreasingBy___closed__1 = _init_l_Lean_Parser_Termination_decreasingBy___closed__1(); -lean_mark_persistent(l_Lean_Parser_Termination_decreasingBy___closed__1); -l_Lean_Parser_Termination_decreasingBy___closed__2 = _init_l_Lean_Parser_Termination_decreasingBy___closed__2(); -lean_mark_persistent(l_Lean_Parser_Termination_decreasingBy___closed__2); -l_Lean_Parser_Termination_decreasingBy___closed__3 = _init_l_Lean_Parser_Termination_decreasingBy___closed__3(); -lean_mark_persistent(l_Lean_Parser_Termination_decreasingBy___closed__3); -l_Lean_Parser_Termination_decreasingBy___closed__4 = _init_l_Lean_Parser_Termination_decreasingBy___closed__4(); -lean_mark_persistent(l_Lean_Parser_Termination_decreasingBy___closed__4); -l_Lean_Parser_Termination_decreasingBy___closed__5 = _init_l_Lean_Parser_Termination_decreasingBy___closed__5(); -lean_mark_persistent(l_Lean_Parser_Termination_decreasingBy___closed__5); -l_Lean_Parser_Termination_decreasingBy___closed__6 = _init_l_Lean_Parser_Termination_decreasingBy___closed__6(); -lean_mark_persistent(l_Lean_Parser_Termination_decreasingBy___closed__6); -l_Lean_Parser_Termination_decreasingBy___closed__7 = _init_l_Lean_Parser_Termination_decreasingBy___closed__7(); -lean_mark_persistent(l_Lean_Parser_Termination_decreasingBy___closed__7); -l_Lean_Parser_Termination_decreasingBy___closed__8 = _init_l_Lean_Parser_Termination_decreasingBy___closed__8(); -lean_mark_persistent(l_Lean_Parser_Termination_decreasingBy___closed__8); -l_Lean_Parser_Termination_decreasingBy___closed__9 = _init_l_Lean_Parser_Termination_decreasingBy___closed__9(); -lean_mark_persistent(l_Lean_Parser_Termination_decreasingBy___closed__9); -l_Lean_Parser_Termination_decreasingBy___closed__10 = _init_l_Lean_Parser_Termination_decreasingBy___closed__10(); -lean_mark_persistent(l_Lean_Parser_Termination_decreasingBy___closed__10); -l_Lean_Parser_Termination_decreasingBy = _init_l_Lean_Parser_Termination_decreasingBy(); -lean_mark_persistent(l_Lean_Parser_Termination_decreasingBy); -l_Lean_Parser_Termination_suffix___closed__1 = _init_l_Lean_Parser_Termination_suffix___closed__1(); -lean_mark_persistent(l_Lean_Parser_Termination_suffix___closed__1); -l_Lean_Parser_Termination_suffix___closed__2 = _init_l_Lean_Parser_Termination_suffix___closed__2(); -lean_mark_persistent(l_Lean_Parser_Termination_suffix___closed__2); -l_Lean_Parser_Termination_suffix___closed__3 = _init_l_Lean_Parser_Termination_suffix___closed__3(); -lean_mark_persistent(l_Lean_Parser_Termination_suffix___closed__3); -l_Lean_Parser_Termination_suffix___closed__4 = _init_l_Lean_Parser_Termination_suffix___closed__4(); -lean_mark_persistent(l_Lean_Parser_Termination_suffix___closed__4); -l_Lean_Parser_Termination_suffix___closed__5 = _init_l_Lean_Parser_Termination_suffix___closed__5(); -lean_mark_persistent(l_Lean_Parser_Termination_suffix___closed__5); -l_Lean_Parser_Termination_suffix___closed__6 = _init_l_Lean_Parser_Termination_suffix___closed__6(); -lean_mark_persistent(l_Lean_Parser_Termination_suffix___closed__6); -l_Lean_Parser_Termination_suffix___closed__7 = _init_l_Lean_Parser_Termination_suffix___closed__7(); -lean_mark_persistent(l_Lean_Parser_Termination_suffix___closed__7); -l_Lean_Parser_Termination_suffix___closed__8 = _init_l_Lean_Parser_Termination_suffix___closed__8(); -lean_mark_persistent(l_Lean_Parser_Termination_suffix___closed__8); -l_Lean_Parser_Termination_suffix___closed__9 = _init_l_Lean_Parser_Termination_suffix___closed__9(); -lean_mark_persistent(l_Lean_Parser_Termination_suffix___closed__9); -l_Lean_Parser_Termination_suffix___closed__10 = _init_l_Lean_Parser_Termination_suffix___closed__10(); -lean_mark_persistent(l_Lean_Parser_Termination_suffix___closed__10); -l_Lean_Parser_Termination_suffix___closed__11 = _init_l_Lean_Parser_Termination_suffix___closed__11(); -lean_mark_persistent(l_Lean_Parser_Termination_suffix___closed__11); -l_Lean_Parser_Termination_suffix = _init_l_Lean_Parser_Termination_suffix(); -lean_mark_persistent(l_Lean_Parser_Termination_suffix); l_Lean_Parser_Term_byTactic___closed__1 = _init_l_Lean_Parser_Term_byTactic___closed__1(); lean_mark_persistent(l_Lean_Parser_Term_byTactic___closed__1); l_Lean_Parser_Term_byTactic___closed__2 = _init_l_Lean_Parser_Term_byTactic___closed__2(); @@ -70848,6 +70668,10 @@ l_Lean_Parser_Term_hole___closed__5 = _init_l_Lean_Parser_Term_hole___closed__5( lean_mark_persistent(l_Lean_Parser_Term_hole___closed__5); l_Lean_Parser_Term_hole___closed__6 = _init_l_Lean_Parser_Term_hole___closed__6(); lean_mark_persistent(l_Lean_Parser_Term_hole___closed__6); +l_Lean_Parser_Term_hole___closed__7 = _init_l_Lean_Parser_Term_hole___closed__7(); +lean_mark_persistent(l_Lean_Parser_Term_hole___closed__7); +l_Lean_Parser_Term_hole___closed__8 = _init_l_Lean_Parser_Term_hole___closed__8(); +lean_mark_persistent(l_Lean_Parser_Term_hole___closed__8); l_Lean_Parser_Term_hole = _init_l_Lean_Parser_Term_hole(); lean_mark_persistent(l_Lean_Parser_Term_hole); if (builtin) {res = l___regBuiltin_Lean_Parser_Term_hole(lean_io_mk_world()); @@ -71251,6 +71075,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); if (builtin) {res = l___regBuiltin_Lean_Parser_Term_typeAscription(lean_io_mk_world()); @@ -75911,6 +75737,102 @@ l_Lean_Parser_Term_attributes___closed__14 = _init_l_Lean_Parser_Term_attributes 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_Termination_terminationBy___closed__1 = _init_l_Lean_Parser_Termination_terminationBy___closed__1(); +lean_mark_persistent(l_Lean_Parser_Termination_terminationBy___closed__1); +l_Lean_Parser_Termination_terminationBy___closed__2 = _init_l_Lean_Parser_Termination_terminationBy___closed__2(); +lean_mark_persistent(l_Lean_Parser_Termination_terminationBy___closed__2); +l_Lean_Parser_Termination_terminationBy___closed__3 = _init_l_Lean_Parser_Termination_terminationBy___closed__3(); +lean_mark_persistent(l_Lean_Parser_Termination_terminationBy___closed__3); +l_Lean_Parser_Termination_terminationBy___closed__4 = _init_l_Lean_Parser_Termination_terminationBy___closed__4(); +lean_mark_persistent(l_Lean_Parser_Termination_terminationBy___closed__4); +l_Lean_Parser_Termination_terminationBy___closed__5 = _init_l_Lean_Parser_Termination_terminationBy___closed__5(); +lean_mark_persistent(l_Lean_Parser_Termination_terminationBy___closed__5); +l_Lean_Parser_Termination_terminationBy___closed__6 = _init_l_Lean_Parser_Termination_terminationBy___closed__6(); +lean_mark_persistent(l_Lean_Parser_Termination_terminationBy___closed__6); +l_Lean_Parser_Termination_terminationBy___closed__7 = _init_l_Lean_Parser_Termination_terminationBy___closed__7(); +lean_mark_persistent(l_Lean_Parser_Termination_terminationBy___closed__7); +l_Lean_Parser_Termination_terminationBy___closed__8 = _init_l_Lean_Parser_Termination_terminationBy___closed__8(); +lean_mark_persistent(l_Lean_Parser_Termination_terminationBy___closed__8); +l_Lean_Parser_Termination_terminationBy___closed__9 = _init_l_Lean_Parser_Termination_terminationBy___closed__9(); +lean_mark_persistent(l_Lean_Parser_Termination_terminationBy___closed__9); +l_Lean_Parser_Termination_terminationBy___closed__10 = _init_l_Lean_Parser_Termination_terminationBy___closed__10(); +lean_mark_persistent(l_Lean_Parser_Termination_terminationBy___closed__10); +l_Lean_Parser_Termination_terminationBy___closed__11 = _init_l_Lean_Parser_Termination_terminationBy___closed__11(); +lean_mark_persistent(l_Lean_Parser_Termination_terminationBy___closed__11); +l_Lean_Parser_Termination_terminationBy___closed__12 = _init_l_Lean_Parser_Termination_terminationBy___closed__12(); +lean_mark_persistent(l_Lean_Parser_Termination_terminationBy___closed__12); +l_Lean_Parser_Termination_terminationBy___closed__13 = _init_l_Lean_Parser_Termination_terminationBy___closed__13(); +lean_mark_persistent(l_Lean_Parser_Termination_terminationBy___closed__13); +l_Lean_Parser_Termination_terminationBy___closed__14 = _init_l_Lean_Parser_Termination_terminationBy___closed__14(); +lean_mark_persistent(l_Lean_Parser_Termination_terminationBy___closed__14); +l_Lean_Parser_Termination_terminationBy___closed__15 = _init_l_Lean_Parser_Termination_terminationBy___closed__15(); +lean_mark_persistent(l_Lean_Parser_Termination_terminationBy___closed__15); +l_Lean_Parser_Termination_terminationBy = _init_l_Lean_Parser_Termination_terminationBy(); +lean_mark_persistent(l_Lean_Parser_Termination_terminationBy); +l_Lean_Parser_Termination_terminationBy_x3f___closed__1 = _init_l_Lean_Parser_Termination_terminationBy_x3f___closed__1(); +lean_mark_persistent(l_Lean_Parser_Termination_terminationBy_x3f___closed__1); +l_Lean_Parser_Termination_terminationBy_x3f___closed__2 = _init_l_Lean_Parser_Termination_terminationBy_x3f___closed__2(); +lean_mark_persistent(l_Lean_Parser_Termination_terminationBy_x3f___closed__2); +l_Lean_Parser_Termination_terminationBy_x3f___closed__3 = _init_l_Lean_Parser_Termination_terminationBy_x3f___closed__3(); +lean_mark_persistent(l_Lean_Parser_Termination_terminationBy_x3f___closed__3); +l_Lean_Parser_Termination_terminationBy_x3f___closed__4 = _init_l_Lean_Parser_Termination_terminationBy_x3f___closed__4(); +lean_mark_persistent(l_Lean_Parser_Termination_terminationBy_x3f___closed__4); +l_Lean_Parser_Termination_terminationBy_x3f___closed__5 = _init_l_Lean_Parser_Termination_terminationBy_x3f___closed__5(); +lean_mark_persistent(l_Lean_Parser_Termination_terminationBy_x3f___closed__5); +l_Lean_Parser_Termination_terminationBy_x3f___closed__6 = _init_l_Lean_Parser_Termination_terminationBy_x3f___closed__6(); +lean_mark_persistent(l_Lean_Parser_Termination_terminationBy_x3f___closed__6); +l_Lean_Parser_Termination_terminationBy_x3f___closed__7 = _init_l_Lean_Parser_Termination_terminationBy_x3f___closed__7(); +lean_mark_persistent(l_Lean_Parser_Termination_terminationBy_x3f___closed__7); +l_Lean_Parser_Termination_terminationBy_x3f___closed__8 = _init_l_Lean_Parser_Termination_terminationBy_x3f___closed__8(); +lean_mark_persistent(l_Lean_Parser_Termination_terminationBy_x3f___closed__8); +l_Lean_Parser_Termination_terminationBy_x3f = _init_l_Lean_Parser_Termination_terminationBy_x3f(); +lean_mark_persistent(l_Lean_Parser_Termination_terminationBy_x3f); +l_Lean_Parser_Termination_decreasingBy___closed__1 = _init_l_Lean_Parser_Termination_decreasingBy___closed__1(); +lean_mark_persistent(l_Lean_Parser_Termination_decreasingBy___closed__1); +l_Lean_Parser_Termination_decreasingBy___closed__2 = _init_l_Lean_Parser_Termination_decreasingBy___closed__2(); +lean_mark_persistent(l_Lean_Parser_Termination_decreasingBy___closed__2); +l_Lean_Parser_Termination_decreasingBy___closed__3 = _init_l_Lean_Parser_Termination_decreasingBy___closed__3(); +lean_mark_persistent(l_Lean_Parser_Termination_decreasingBy___closed__3); +l_Lean_Parser_Termination_decreasingBy___closed__4 = _init_l_Lean_Parser_Termination_decreasingBy___closed__4(); +lean_mark_persistent(l_Lean_Parser_Termination_decreasingBy___closed__4); +l_Lean_Parser_Termination_decreasingBy___closed__5 = _init_l_Lean_Parser_Termination_decreasingBy___closed__5(); +lean_mark_persistent(l_Lean_Parser_Termination_decreasingBy___closed__5); +l_Lean_Parser_Termination_decreasingBy___closed__6 = _init_l_Lean_Parser_Termination_decreasingBy___closed__6(); +lean_mark_persistent(l_Lean_Parser_Termination_decreasingBy___closed__6); +l_Lean_Parser_Termination_decreasingBy___closed__7 = _init_l_Lean_Parser_Termination_decreasingBy___closed__7(); +lean_mark_persistent(l_Lean_Parser_Termination_decreasingBy___closed__7); +l_Lean_Parser_Termination_decreasingBy___closed__8 = _init_l_Lean_Parser_Termination_decreasingBy___closed__8(); +lean_mark_persistent(l_Lean_Parser_Termination_decreasingBy___closed__8); +l_Lean_Parser_Termination_decreasingBy___closed__9 = _init_l_Lean_Parser_Termination_decreasingBy___closed__9(); +lean_mark_persistent(l_Lean_Parser_Termination_decreasingBy___closed__9); +l_Lean_Parser_Termination_decreasingBy___closed__10 = _init_l_Lean_Parser_Termination_decreasingBy___closed__10(); +lean_mark_persistent(l_Lean_Parser_Termination_decreasingBy___closed__10); +l_Lean_Parser_Termination_decreasingBy = _init_l_Lean_Parser_Termination_decreasingBy(); +lean_mark_persistent(l_Lean_Parser_Termination_decreasingBy); +l_Lean_Parser_Termination_suffix___closed__1 = _init_l_Lean_Parser_Termination_suffix___closed__1(); +lean_mark_persistent(l_Lean_Parser_Termination_suffix___closed__1); +l_Lean_Parser_Termination_suffix___closed__2 = _init_l_Lean_Parser_Termination_suffix___closed__2(); +lean_mark_persistent(l_Lean_Parser_Termination_suffix___closed__2); +l_Lean_Parser_Termination_suffix___closed__3 = _init_l_Lean_Parser_Termination_suffix___closed__3(); +lean_mark_persistent(l_Lean_Parser_Termination_suffix___closed__3); +l_Lean_Parser_Termination_suffix___closed__4 = _init_l_Lean_Parser_Termination_suffix___closed__4(); +lean_mark_persistent(l_Lean_Parser_Termination_suffix___closed__4); +l_Lean_Parser_Termination_suffix___closed__5 = _init_l_Lean_Parser_Termination_suffix___closed__5(); +lean_mark_persistent(l_Lean_Parser_Termination_suffix___closed__5); +l_Lean_Parser_Termination_suffix___closed__6 = _init_l_Lean_Parser_Termination_suffix___closed__6(); +lean_mark_persistent(l_Lean_Parser_Termination_suffix___closed__6); +l_Lean_Parser_Termination_suffix___closed__7 = _init_l_Lean_Parser_Termination_suffix___closed__7(); +lean_mark_persistent(l_Lean_Parser_Termination_suffix___closed__7); +l_Lean_Parser_Termination_suffix___closed__8 = _init_l_Lean_Parser_Termination_suffix___closed__8(); +lean_mark_persistent(l_Lean_Parser_Termination_suffix___closed__8); +l_Lean_Parser_Termination_suffix___closed__9 = _init_l_Lean_Parser_Termination_suffix___closed__9(); +lean_mark_persistent(l_Lean_Parser_Termination_suffix___closed__9); +l_Lean_Parser_Termination_suffix___closed__10 = _init_l_Lean_Parser_Termination_suffix___closed__10(); +lean_mark_persistent(l_Lean_Parser_Termination_suffix___closed__10); +l_Lean_Parser_Termination_suffix___closed__11 = _init_l_Lean_Parser_Termination_suffix___closed__11(); +lean_mark_persistent(l_Lean_Parser_Termination_suffix___closed__11); +l_Lean_Parser_Termination_suffix = _init_l_Lean_Parser_Termination_suffix(); +lean_mark_persistent(l_Lean_Parser_Termination_suffix); l_Lean_Parser_Term_letRecDecl___closed__1 = _init_l_Lean_Parser_Term_letRecDecl___closed__1(); lean_mark_persistent(l_Lean_Parser_Term_letRecDecl___closed__1); l_Lean_Parser_Term_letRecDecl___closed__2 = _init_l_Lean_Parser_Term_letRecDecl___closed__2(); @@ -76138,10 +76060,6 @@ l_Lean_Parser_Termination_terminationBy_formatter___closed__9 = _init_l_Lean_Par lean_mark_persistent(l_Lean_Parser_Termination_terminationBy_formatter___closed__9); l_Lean_Parser_Termination_terminationBy_formatter___closed__10 = _init_l_Lean_Parser_Termination_terminationBy_formatter___closed__10(); lean_mark_persistent(l_Lean_Parser_Termination_terminationBy_formatter___closed__10); -l_Lean_Parser_Termination_terminationBy_formatter___closed__11 = _init_l_Lean_Parser_Termination_terminationBy_formatter___closed__11(); -lean_mark_persistent(l_Lean_Parser_Termination_terminationBy_formatter___closed__11); -l_Lean_Parser_Termination_terminationBy_formatter___closed__12 = _init_l_Lean_Parser_Termination_terminationBy_formatter___closed__12(); -lean_mark_persistent(l_Lean_Parser_Termination_terminationBy_formatter___closed__12); l___regBuiltin_Lean_Parser_Termination_terminationBy_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Termination_terminationBy_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Termination_terminationBy_formatter___closed__1); l___regBuiltin_Lean_Parser_Termination_terminationBy_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Termination_terminationBy_formatter___closed__2(); @@ -76381,10 +76299,6 @@ l_Lean_Parser_Termination_terminationBy_parenthesizer___closed__9 = _init_l_Lean lean_mark_persistent(l_Lean_Parser_Termination_terminationBy_parenthesizer___closed__9); l_Lean_Parser_Termination_terminationBy_parenthesizer___closed__10 = _init_l_Lean_Parser_Termination_terminationBy_parenthesizer___closed__10(); lean_mark_persistent(l_Lean_Parser_Termination_terminationBy_parenthesizer___closed__10); -l_Lean_Parser_Termination_terminationBy_parenthesizer___closed__11 = _init_l_Lean_Parser_Termination_terminationBy_parenthesizer___closed__11(); -lean_mark_persistent(l_Lean_Parser_Termination_terminationBy_parenthesizer___closed__11); -l_Lean_Parser_Termination_terminationBy_parenthesizer___closed__12 = _init_l_Lean_Parser_Termination_terminationBy_parenthesizer___closed__12(); -lean_mark_persistent(l_Lean_Parser_Termination_terminationBy_parenthesizer___closed__12); l___regBuiltin_Lean_Parser_Termination_terminationBy_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Termination_terminationBy_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Termination_terminationBy_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Termination_terminationBy_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Termination_terminationBy_parenthesizer___closed__2(); @@ -80507,117 +80421,117 @@ lean_mark_persistent(l___regBuiltin_Lean_Parser_Tactic_quotSeq_parenthesizer___c if (builtin) {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_5127____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__1(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__1); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__2(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__2); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__3 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__3(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__3); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__4 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__4(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__4); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__5 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__5(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__5); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__6 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__6(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__6); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__7 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__7(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__7); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__8 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__8(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__8); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__9 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__9(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__9); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__10 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__10(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__10); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__11 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__11(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__11); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__12 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__12(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__12); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__13 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__13(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__13); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__14 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__14(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__14); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__15 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__15(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__15); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__16 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__16(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__16); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__17 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__17(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__17); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__18 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__18(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__18); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__19 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__19(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__19); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__20 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__20(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__20); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__21 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__21(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__21); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__22 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__22(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__22); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__23 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__23(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__23); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__24 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__24(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__24); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__25 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__25(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__25); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__26 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__26(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__26); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__27 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__27(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__27); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__28 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__28(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__28); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__29 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__29(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__29); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__30 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__30(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__30); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__31 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__31(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__31); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__32 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__32(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__32); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__33 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__33(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__33); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__34 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__34(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__34); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__35 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__35(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__35); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__36 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__36(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__36); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__37 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__37(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__37); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__38 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__38(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__38); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__39 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__39(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__39); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__40 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__40(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__40); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__41 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__41(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__41); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__42 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__42(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__42); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__43 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__43(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__43); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__44 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__44(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__44); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__45 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__45(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__45); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__46 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__46(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__46); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__47 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__47(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__47); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__48 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__48(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__48); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__49 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__49(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__49); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__50 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__50(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__50); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__51 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__51(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__51); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__52 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__52(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__52); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__53 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__53(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__53); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__54 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__54(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__54); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__55 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__55(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127____closed__55); -if (builtin) {res = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5127_(lean_io_mk_world()); +}l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__1(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__1); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__2(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__2); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__3 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__3(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__3); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__4 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__4(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__4); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__5 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__5(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__5); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__6 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__6(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__6); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__7 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__7(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__7); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__8 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__8(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__8); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__9 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__9(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__9); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__10 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__10(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__10); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__11 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__11(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__11); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__12 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__12(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__12); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__13 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__13(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__13); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__14 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__14(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__14); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__15 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__15(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__15); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__16 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__16(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__16); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__17 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__17(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__17); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__18 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__18(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__18); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__19 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__19(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__19); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__20 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__20(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__20); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__21 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__21(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__21); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__22 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__22(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__22); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__23 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__23(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__23); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__24 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__24(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__24); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__25 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__25(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__25); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__26 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__26(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__26); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__27 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__27(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__27); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__28 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__28(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__28); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__29 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__29(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__29); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__30 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__30(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__30); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__31 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__31(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__31); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__32 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__32(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__32); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__33 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__33(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__33); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__34 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__34(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__34); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__35 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__35(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__35); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__36 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__36(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__36); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__37 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__37(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__37); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__38 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__38(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__38); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__39 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__39(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__39); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__40 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__40(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__40); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__41 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__41(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__41); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__42 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__42(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__42); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__43 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__43(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__43); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__44 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__44(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__44); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__45 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__45(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__45); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__46 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__46(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__46); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__47 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__47(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__47); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__48 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__48(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__48); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__49 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__49(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__49); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__50 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__50(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__50); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__51 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__51(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__51); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__52 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__52(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__52); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__53 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__53(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__53); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__54 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__54(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__54); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__55 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__55(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126____closed__55); +if (builtin) {res = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5126_(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.c b/stage0/stdlib/Lean/PrettyPrinter.c index b81c2924135b..9977087aa72e 100644 --- a/stage0/stdlib/Lean/PrettyPrinter.c +++ b/stage0/stdlib/Lean/PrettyPrinter.c @@ -87,6 +87,7 @@ static lean_object* l_Lean_PrettyPrinter_ppExprLegacy___closed__3; static lean_object* l_Lean_PrettyPrinter_registerParserCompilers___closed__8; static lean_object* l_Lean_PPContext_runCoreM___rarg___closed__6; lean_object* lean_st_ref_get(lean_object*, lean_object*); +lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_ppConst___lambda__1(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_Name_num___override(lean_object*, lean_object*); @@ -141,7 +142,6 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_ppModule(lean_object*, lean_object static lean_object* l_Lean_PPContext_runMetaM___rarg___closed__15; lean_object* l_Lean_PrettyPrinter_parenthesizeCategory(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_ppConst___lambda__1___closed__1; -lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature(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_786____lambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_838____closed__5; static lean_object* l_Lean_PrettyPrinter_registerParserCompilers___closed__3; @@ -1618,9 +1618,12 @@ return x_14; static lean_object* _init_l_Lean_PrettyPrinter_ppSignature___closed__1() { _start: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature), 7, 0); -return x_1; +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_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___boxed), 8, 1); +lean_closure_set(x_3, 0, x_2); +return x_3; } } LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_ppSignature(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { diff --git a/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Builtins.c b/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Builtins.c index 34e1be770062..ac12ec151eca 100644 --- a/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Builtins.c +++ b/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Builtins.c @@ -13,11 +13,9 @@ #ifdef __cplusplus extern "C" { #endif -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabSort___closed__10; LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__41(lean_object*, size_t, size_t, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabFVar___closed__5; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda__3___closed__1; LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -43,12 +41,12 @@ LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppImp lean_object* l_Lean_Meta_getCoeFnInfo_x3f(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_getPPStructureInstanceType___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__8(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabNeg(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*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabMVar___closed__4; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_useAppExplicit___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_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__12(lean_object*, size_t, size_t, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabFVar___closed__8; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabDIte(lean_object*); @@ -63,13 +61,15 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabOfNa LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___spec__5___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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_PrettyPrinter_Delaborator_delabOfNatCore(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_mkNameLit(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_getParamKinds___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_firstM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryUnexpand___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_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNames(lean_object*); +LEAN_EXPORT lean_object* l_Lean_ofExcept___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_RBNode_insert___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_annotateBoolAt___spec__2(lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabOfNat___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withMDataExpr___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__2(lean_object*); @@ -90,7 +90,6 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_ParamKind_isRegularExp LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabLit___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabForall___lambda__1___closed__3; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabSigmaCore___lambda__1___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabPSigma(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabProjectionApp___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -104,9 +103,7 @@ static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabCond___cl LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabAppMatch(lean_object*); static lean_object* l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2___closed__3; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabCond___lambda__1___closed__3; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__8(lean_object*, size_t, size_t, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabApp___closed__4; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__38___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabProj___closed__2; lean_object* lean_private_to_user_name(lean_object*); lean_object* l_Lean_PrettyPrinter_Formatter_orelse_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -114,26 +111,25 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetFun___closed__3; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabNameMkNum___closed__2; static lean_object* l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__11; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__30(lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabSort___closed__8; lean_object* l_Lean_getPPNotation___boxed(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__6; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_foldrMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabForall___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_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConst___spec__6___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingDomain___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__16(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___closed__5; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_delabLam___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_pp_fullNames; -LEAN_EXPORT lean_object* l_Lean_ofExcept___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__22___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__15___closed__2; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda__8___closed__7; LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabProj___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_delabLam___lambda__3___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabApp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabCond___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__18(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_instInhabitedParamKind; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__3; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__13(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_whnf(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__3___closed__1; @@ -144,12 +140,11 @@ lean_object* l_Lean_Core_instantiateTypeLevelParams(lean_object*, lean_object*, static lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__3___closed__12; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__3; lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_getPos___at_Lean_PrettyPrinter_Delaborator_getOptionsAtCurrPos___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabMVar___closed__5; -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__18___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_isProp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabLam___closed__4; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_delabConst___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_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__24___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_shouldGroupWithNext___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_getPPFullNames___boxed(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_declSigWithId_formatter___closed__1; @@ -157,7 +152,6 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabOfNat___closed__3; lean_object* l_Lean_getPPFunBinderTypes___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_shouldGroupWithNext(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabMData___spec__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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda__4___closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withType___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabMData___closed__5; @@ -166,7 +160,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBod static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__9; static lean_object* l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__8; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConst___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__33___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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_PrettyPrinter_Delaborator_reifyName___closed__1; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_declSigWithId_formatter___closed__4; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabFVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -176,15 +170,15 @@ LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at_Lean_PrettyPrinter_Del static lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabOfNat___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDIte(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__28(lean_object*, lean_object*, size_t, size_t, lean_object*); -static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___closed__1; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__12; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabForall___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabSigma(lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabChar___closed__2; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDivRatCore___closed__2; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__20___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_ofExcept___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getAutoParamTactic_x3f(lean_object*); uint8_t l_Lean_Expr_isLet(lean_object*); lean_object* l_Lean_PrettyPrinter_Delaborator_getPPOption___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -198,12 +192,10 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabHDiv___closed__1; uint8_t l_Lean_Expr_isApp(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_hasIdent___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_toString(lean_object*, uint8_t); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_leadingNode(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withMDataExpr___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabMData___spec__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getId(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__21___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDivRatCore___closed__6; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLit___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__2; @@ -212,25 +204,24 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabOfNatCore___boxed lean_object* l_Lean_Parser_Command_declSig_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkNamedArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabOfScientific___closed__4; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_declSigWithId_formatter___closed__3; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_loop___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__1___closed__1; lean_object* l_Lean_PrettyPrinter_Delaborator_instInhabitedDelabM(lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__12(lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda__4___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__3; LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabProj___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabSigma___closed__1; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabMVar___closed__1; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__40(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabProjectionApp_withoutParentProjections___spec__6(lean_object*); -static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__14; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabChar___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__7___boxed(lean_object**); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabSigmaCore___lambda__1___closed__7; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_ofExcept___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__22(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Array_contains___at_Lean_findField_x3f___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLam___lambda__1___closed__2; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabProj___closed__6; @@ -238,7 +229,9 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDivRatCore___boxe LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabProjectionApp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_withMDatasOptions(lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabForall___closed__3; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabLit(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda__3___closed__2; uint8_t lean_usize_dec_eq(size_t, size_t); @@ -248,7 +241,6 @@ extern uint8_t l_instInhabitedBool; LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_getPPFunBinderTypes(lean_object*); -static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__8___closed__1; lean_object* l_Lean_Syntax_getArgs(lean_object*); static lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__3___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConst___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -259,6 +251,7 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConst___closed__13; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConst___closed__9; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabNegIntCore___closed__2; static lean_object* l_Lean_PrettyPrinter_Delaborator_maybeAddBlockImplicit___closed__4; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__14(lean_object*, lean_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_delabForall___lambda__2___closed__2; lean_object* lean_array_fget(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabOfScientific___lambda__1___closed__4; @@ -269,14 +262,14 @@ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delabora static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda__4___closed__2; LEAN_EXPORT lean_object* l_List_firstM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryUnexpand___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_panic___at_Lean_PrettyPrinter_Delaborator_delabProj___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__13(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLam___lambda__1___closed__4; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_AppMatchState_varNames___default; -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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*); uint8_t l_Lean_Expr_isLetFun(lean_object*); lean_object* l_ReaderT_instMonadReaderT___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__28___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_fvarId_x21(lean_object*); lean_object* l_Lean_PrettyPrinter_Delaborator_isNonConstFun(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__4___closed__2; @@ -287,6 +280,7 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabApp_delabAppFn___close lean_object* l_ReaderT_pure___at_Lean_PrettyPrinter_Delaborator_instMonadQuotationDelabM___spec__1___rarg___boxed(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*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__12; +lean_object* l_Lean_RBNode_insert___at_Lean_NameSet_insert___spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabFVar___closed__6; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabProjectionApp_withoutParentProjections___spec__2(lean_object*); lean_object* l_Lean_getPPStructureProjections___boxed(lean_object*); @@ -295,41 +289,35 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_L lean_object* l_Lean_Expr_getBoundedAppFn(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withMDataExpr___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__2___rarg___closed__2; static lean_object* l_Lean_PrettyPrinter_Delaborator_declSigWithId___closed__4; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__9(lean_object*, size_t, size_t, 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*); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__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_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabOfScientific(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_PrettyPrinter_Delaborator_delabLetE___spec__1(lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabApp___closed__5; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_useAppExplicit(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_declSigWithId_parenthesizer___closed__1; lean_object* l_List_head_x3f___rarg(lean_object*); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__15(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__34___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___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_PrettyPrinter_Delaborator_SubExpr_nextExtraPos___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___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_panic___at_Lean_PrettyPrinter_Delaborator_delabMData___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_delabConstWithSignature_delabParamsAux___lambda__18___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__11___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_PrettyPrinter_Delaborator_withMDataOptions___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_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__10(lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabMData___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_delabMData___closed__3; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_PrettyPrinter_Delaborator_delabLetE___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_useAppExplicit___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_Syntax_node5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__19(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_delabConstWithSignature_delabTy___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabSort(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_Meta_forallTelescopeReducing___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___spec__6(lean_object*); lean_object* l_Lean_Name_mkStr5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__36(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabLam___spec__2(lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_stringToMessageData(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetE(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -338,16 +326,15 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabLam___lambda__3(l lean_object* l_Lean_Parser_termParser_formatter___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLam___lambda__1___closed__9; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabProjectionApp___closed__1; +static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___closed__3; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__3___closed__7; LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___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_panic___at_Lean_PrettyPrinter_Delaborator_delabProj___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_getParamKinds___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabForall(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__17(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabLam___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabMVar(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__42(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_useAppExplicit___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__11; LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda(lean_object*); @@ -368,23 +355,25 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConst(lean_object static lean_object* l_Lean_PrettyPrinter_Delaborator_delabMVar___closed__1; static lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__1; uint8_t l_Lean_LocalContext_usesUserName(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__18___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabSort(lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabLit___closed__4; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabNameMkNum___closed__4; static lean_object* l_Lean_PrettyPrinter_Delaborator_declSigWithId___closed__6; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabForall___lambda__1___closed__1; lean_object* l_Lean_getPPUniverses___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__27___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_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__18(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___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_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__30___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_evalSyntaxConstantUnsafe___closed__2; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabProj___closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_shouldShowMotive(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_Lean_PrettyPrinter_Delaborator_delabOfScientific___lambda__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConst___closed__8; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__7; @@ -392,9 +381,8 @@ LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__ LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabNegIntCore___closed__5; static lean_object* l_Lean_PrettyPrinter_Delaborator_maybeAddBlockImplicit___closed__2; -static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__1___closed__1; extern lean_object* l_instInhabitedNat; -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getNumArgs(lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabChar___closed__5; lean_object* l_Lean_getPPMatch___boxed(lean_object*); @@ -405,12 +393,12 @@ LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_withMDataOp LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabBinders(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabForallBinders___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_PrettyPrinter_Delaborator_delabAppImplicitCore_mkNamedArg(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_delabConstWithSignature_delabParams___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*); extern lean_object* l_Lean_projectionFnInfoExt; lean_object* l_Lean_Syntax_mkApp(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabSigmaCore___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__32___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabNeg___closed__1; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__17(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabProjectionApp_withoutParentProjections___spec__7(lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabNeg___closed__4; lean_object* l_Lean_isLHSGoal_x3f(lean_object*); @@ -422,8 +410,6 @@ static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabSort___cl LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabFVar___closed__1; uint8_t l_Lean_Expr_isLambda(lean_object*); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__19___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__13(lean_object*, size_t, size_t, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabLetFun___closed__5; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLam___lambda__3___closed__7; static lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__2___closed__4; @@ -435,35 +421,33 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Delab static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_coeDelaborator___closed__3; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDIte___closed__2; lean_object* l_Lean_getPPPrivateNames___boxed(lean_object*); +static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___closed__4; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConst___closed__1; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabHDiv___closed__1; 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_object* lean_expr_instantiate_rev_range(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__23(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLam___lambda__3___closed__6; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabMData___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___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_PrettyPrinter_Delaborator_declSigWithId___closed__2; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__39___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__7; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__1___closed__2; size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_useAppExplicit___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_isSubobjectField_x3f(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetE___closed__2; -static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__16___closed__1; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_getParamKinds___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_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__10; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_loop___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__1___closed__2; static lean_object* l_Lean_PrettyPrinter_Delaborator_instInhabitedParamKind___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_declSigWithId_parenthesizer___closed__3; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLam___lambda__3___closed__4; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabOfScientific___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_lambdaTelescopeImp___rarg(lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__2___boxed(lean_object**); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__34(lean_object*, lean_object*, size_t, size_t, lean_object*); extern lean_object* l_Lean_Parser_Command_declSig; static lean_object* l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__14; @@ -476,6 +460,7 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delabora LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabLit(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___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_PrettyPrinter_Delaborator_delabDoElems___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2___closed__4; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_withLocalDecl___at_Lean_PrettyPrinter_Delaborator_withBindingBodyUnusedName___spec__3___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*); @@ -496,8 +481,6 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabProjectionApp___closed LEAN_EXPORT lean_object* l_Lean_resolveGlobalName___at_Lean_PrettyPrinter_Delaborator_delabConst___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabDo___spec__1___closed__1; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__16(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabOfNat(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_pp_piBinderTypes; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabMData___closed__3; @@ -536,18 +519,18 @@ lean_object* l___private_Lean_Expr_0__Lean_Expr_getBoundedAppArgsAux(lean_object 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_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__3___closed__2; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabMVar___closed__3; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__32(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabForall___lambda__1(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_PrettyPrinter_Delaborator_SubExpr_withAppFn___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_inaccessible_x3f(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDo___lambda__1___closed__2; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__40___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__17(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__4(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__21(lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabOfScientific___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___spec__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__2(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__4; static lean_object* l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDivRatCore(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -555,29 +538,34 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda_ LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabMData(lean_object*); static lean_object* l_Array_foldrMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabForall___spec__1___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabForall(lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__2(lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* l_Lean_Syntax_mkStrLit(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabMVar___closed__3; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDIte_delabBranch___closed__1; lean_object* l_Lean_getRevAliases(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__7___closed__1; static lean_object* l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2___closed__2; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__29(lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabNameMkStr___lambda__1___closed__3; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__5___closed__3; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__16(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConst___closed__5; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDivRatCore___closed__5; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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___regBuiltin_Lean_PrettyPrinter_Delaborator_delabChar___closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___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*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabLetE___closed__2; LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabMData___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_delabNameMkNum(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_declSigWithId_parenthesizer(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__25___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_getPPMotivesPi(lean_object*); extern lean_object* l_Lean_PrettyPrinter_parenthesizerAttribute; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabSigma___closed__5; extern lean_object* l_Lean_SubExpr_Pos_typeCoord; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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_PrettyPrinter_Delaborator_delabLam___lambda__3___closed__8; static lean_object* l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -591,14 +579,17 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda_ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabSort___closed__7; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_getPPInstanceTypes___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabProjectionApp_projInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PrettyPrinter_Delaborator_AppMatchState_motiveNamed___default; LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabProj(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__33(lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabLetE___closed__5; extern lean_object* l_Lean_Meta_Match_instInhabitedDiscrInfo; +static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__15___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___closed__2; lean_object* l_Lean_FVarId_getDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLam___lambda__3___closed__10; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabSigmaCore___lambda__1___closed__4; @@ -616,14 +607,16 @@ lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l_Array_mapIdxM_map___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___spec__5___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingDomain___at_Lean_PrettyPrinter_Delaborator_delabLam___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLam___lambda__1___closed__1; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__25(lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabOfScientific___lambda__2___closed__2; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__15(lean_object*, lean_object*, 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_delabAppMatch___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConst___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabProjectionApp_withoutParentProjections___spec__3(lean_object*); lean_object* lean_array_pop(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__3___closed__3; lean_object* l_Lean_getPPUnicodeFun___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___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_Delaborator_delabDoElems___lambda__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_Lean_unresolveNameGlobal___at_Lean_PrettyPrinter_Delaborator_delabConst___spec__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -640,20 +633,20 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delabora LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabForallBinders___lambda__1(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___regBuiltin_Lean_PrettyPrinter_Delaborator_delabProj___closed__2; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_delabLam___spec__4(lean_object*, lean_object*, size_t, size_t); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___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_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_PrettyPrinter_Delaborator_delabLetE___spec__1___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__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_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingDomain___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__3(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_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_getParamKinds___spec__1___lambda__1___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___closed__4; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabNegIntCore___closed__3; -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__5(lean_object*, lean_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_declSigWithId_formatter___closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabProjectionApp_withoutParentProjections___spec__6___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_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__3; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__13___boxed(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_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__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_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_evalSyntaxConstantUnsafe___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabForall___lambda__2___closed__3; lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); @@ -668,13 +661,11 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lamb static lean_object* l_Lean_PrettyPrinter_Delaborator_delabProjectionApp_projInfo___closed__1; lean_object* l_instInhabited___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__8(lean_object*); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_ParamKind_defVal___default; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__4___closed__2; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabPSigma___closed__1; static lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_getParamKinds___spec__1___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDIte___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_delabConstWithSignature_delabParams___lambda__8___closed__3; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__7; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__4___closed__4; static lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___closed__1; @@ -686,7 +677,6 @@ static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabBVar___cl extern lean_object* l_Lean_levelZero; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_nextExtraPos___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabForallBinders___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_delabConstWithSignature_delabParams___lambda__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_declSigWithId_parenthesizer___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabNegIntCore___closed__4; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -702,6 +692,8 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda__8___ lean_object* l_Lean_PrettyPrinter_Delaborator_withAnnotateTermInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_getParamKinds___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabPSigma(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_mkLit(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_unresolveNameGlobal_unresolveNameCore___at_Lean_PrettyPrinter_Delaborator_delabConst___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_PrettyPrinter_Delaborator_delabDo___lambda__1___closed__3; @@ -711,20 +703,20 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_withTypeAscription___b static lean_object* l_Lean_PrettyPrinter_Delaborator_delabSort___closed__4; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabProjectionApp_withoutParentProjections___spec__8___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_declSigWithId_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_BinderInfo_isInstImplicit(uint8_t); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDIte___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabProjectionApp_withoutParentProjections___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_withAntiquot(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__10(lean_object*, lean_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_PrettyPrinter_Delaborator_delabConst___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*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabNeg___closed__1; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__31(lean_object*, lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__14(lean_object*, size_t, size_t, lean_object*); lean_object* l_List_lengthTRAux___rarg(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_useAppExplicit___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_PrettyPrinter_Delaborator_delabConst___closed__4; -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__10(lean_object*, lean_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_delabConstWithSignature_delabParamsAux___lambda__7___closed__3; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabOfScientific(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__37___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_componentsRev(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__5___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabForall___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*); @@ -732,7 +724,6 @@ uint8_t lean_name_eq(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabBVar___closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___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_Name_str___override(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_getPPCoercions___boxed(lean_object*); extern lean_object* l_Lean_rootNamespace; @@ -742,10 +733,8 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabProj___closed__3; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__13; lean_object* l_Lean_Meta_lambdaTelescope___at_Lean_PrettyPrinter_Delaborator_returnsPi___spec__1___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_getPPTagAppFns___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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*); uint8_t l_Lean_Expr_isArrow(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabBVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetE___closed__3; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabOfNat___closed__1; lean_object* l_Lean_Syntax_node2(lean_object*, lean_object*, lean_object*, lean_object*); @@ -756,7 +745,6 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_getParamKinds(lean_obj LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withMDataExpr___at_Lean_PrettyPrinter_Delaborator_delabMData___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabLetFun___closed__2; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabSigmaCore___lambda__1___closed__2; -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabSigmaCore___lambda__3___closed__1; lean_object* l_Lean_getPPAnalysisHole___boxed(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabSort___closed__11; @@ -768,14 +756,15 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabSigmaCore(uint8_t LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__5___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_PrettyPrinter_Delaborator_AppMatchState_discrs___default; lean_object* l_Lean_RBNode_find___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_annotateBoolAt___spec__5(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__9(lean_object*, lean_object*, size_t, size_t, lean_object*); uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabLam___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__3___closed__14; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabLam___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withMDataExpr___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__2___rarg___closed__3; LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__10___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabProj___closed__5; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabOfNatCore___closed__1; lean_object* l_Lean_PrettyPrinter_Delaborator_annotateTermInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -792,10 +781,10 @@ LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_d LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabHDiv(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__4; static lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__3___closed__13; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__31___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabNameMkNum___closed__3; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabDIte___closed__2; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabLam___closed__2; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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_Expr_appFn_x21(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabNameMkStr___closed__2; lean_object* lean_mk_syntax_ident(lean_object*); @@ -815,6 +804,7 @@ LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabProj__ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore(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_PrettyPrinter_Delaborator_delabSort___closed__6; lean_object* l_Lean_Parser_andthen(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__10(lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withMDataExpr___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__2___rarg___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabMData___closed__2; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__8; @@ -822,12 +812,9 @@ lean_object* l_Lean_Expr_getOptParamDefault_x3f(lean_object*); LEAN_EXPORT uint8_t l_Lean_PrettyPrinter_Delaborator_ParamKind_isRegularExplicit(lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabMVar___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabApp_delabAppFn___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_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__39(lean_object*, size_t, size_t, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabLetFun___closed__3; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDIte___lambda__1___closed__6; -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__17(lean_object*, lean_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_declSigWithId_parenthesizer___closed__5; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__38(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_AppMatchState_motive___default; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabDo___closed__3; @@ -836,7 +823,6 @@ lean_object* l_Lean_getPPAnalysisSkip___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_PrettyPrinter_Delaborator_delabLetE___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_Lean_Meta_lambdaTelescope___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__4(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLit___closed__4; -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabLetFun___closed__4; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_declSigWithId_parenthesizer___closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -850,7 +836,6 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabName static lean_object* l_Lean_PrettyPrinter_Delaborator_reifyName___closed__3; lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withType___at_Lean_PrettyPrinter_Delaborator_delab___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabSort___closed__3; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__35___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabMVar___closed__6; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___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_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__2___closed__2; @@ -858,13 +843,14 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabProj___closed__4; lean_object* l_Lean_PrettyPrinter_Delaborator_annotatePos(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabSigma(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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* l_Lean_SubExpr_Pos_push(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabNamedPattern___closed__1; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabCond___closed__1; uint8_t l_Lean_Expr_binderInfo(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda__3___closed__3; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__15(lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabNameMkNum___closed__1; LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabDo___spec__1___closed__2; @@ -885,15 +871,13 @@ 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*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabProjectionApp___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isAutoParam(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__24(lean_object*, lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__29___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_reifyName___closed__2; static lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabDo___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_Delaborator_delabSort___closed__5; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__36___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConst___closed__6; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabDIte___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__3___rarg(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_delabAppExplicitCore___lambda__5___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__6; @@ -905,9 +889,8 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDivRatCore___closed__4 static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__5; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_declSigWithId_parenthesizer___closed__4; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabOfScientific___lambda__2___closed__1; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__42___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__4; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__14(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__12(lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabOfScientific___closed__3; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetFun___closed__1; @@ -916,7 +899,6 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryApp extern lean_object* l_Lean_PrettyPrinter_Delaborator_delabFailureId; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabHDiv___closed__3; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabForall___lambda__1___closed__4; -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_nextExtraPos___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___closed__3; @@ -941,6 +923,7 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLit___closed__5; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabLam___closed__3; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLam___lambda__1___closed__10; static lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__4___closed__1; +extern lean_object* l_Lean_NameSet_empty; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDivRatCore___closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabApp_delabAppFn___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_delabAppImplicitCore_mkNamedArg___closed__1; @@ -948,7 +931,6 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabSigmaCore___lambda__1_ static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabProj___closed__4; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabOfNat___closed__4; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLam___lambda__1___closed__7; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__27(lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__4___closed__3; lean_object* lean_string_length(lean_object*); @@ -977,6 +959,7 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance__ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___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_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabProjectionApp_withoutParentProjections___spec__5(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabSort___closed__12; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* l_Lean_Syntax_SepArray_ofElems(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDo___lambda__1___closed__1; @@ -987,22 +970,18 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_del static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabAppMatch___closed__2; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__8; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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_PrettyPrinter_Delaborator_delabNegIntCore___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__6(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems_prependAndRec(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_Delaborator_failure___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_nextExtraPos___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDIte_delabBranch___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_PrettyPrinter_Delaborator_delabProjectionApp___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_panic___at_Lean_PrettyPrinter_Delaborator_delabProj___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__35(lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__5(lean_object*, size_t, size_t, lean_object*); lean_object* l_Lean_Parser_withCache(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabLetE___closed__1; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__4___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_foldrMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabForall___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__37(lean_object*, size_t, size_t, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabPSigma___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabProjectionApp(lean_object*); @@ -1021,7 +1000,6 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__4(lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabFVar___closed__4; static lean_object* l_Lean_PrettyPrinter_Delaborator_declSigWithId_formatter___closed__4; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__19(lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabSigmaCore___lambda__1___closed__3; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabMVar___closed__4; LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_evalSyntaxConstantUnsafe(lean_object*, lean_object*, lean_object*); @@ -1041,10 +1019,11 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_hasIdent___boxed(lean_ uint8_t l_Lean_BinderInfo_isExplicit(uint8_t); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_needsExplicit___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabForall___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_throwError___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__23___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__4(lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_reifyName___closed__4; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabProj___closed__1; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_node1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__3___closed__5; lean_object* l_Lean_Parser_Command_declSig_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1052,15 +1031,17 @@ lean_object* l_Lean_mkSepArray(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLam___lambda__3___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabLetFun(lean_object*); uint8_t l_Lean_Expr_isConstOf(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__7(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_PrettyPrinter_Delaborator_unexpandStructureInstance___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_NameSet_contains(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabCond___closed__2; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabChar___closed__6; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_AppMatchState_moreArgs___default; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__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_PrettyPrinter_Delaborator_delabConstWithSignature(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDo___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabDIte___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__4___closed__3; @@ -1069,24 +1050,24 @@ static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabHDiv___cl lean_object* l_Lean_Level_dec(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__6(lean_object*, size_t, size_t, lean_object*); lean_object* l_Lean_LocalDecl_type(lean_object*); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_Array_foldrMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabForall___spec__1___closed__2; lean_object* lean_panic_fn(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__7(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__1___closed__2; lean_object* l_Lean_getPPNumericTypes___boxed(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__3___closed__4; lean_object* l_Lean_getPPInstances___boxed(lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabMData___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__26___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_coeDelaborator(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabChar(lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabBVar___closed__4; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__3(lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabBVar___closed__3; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_delabConst___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* 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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabProj___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_tailD___rarg(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabSigmaCore___lambda__1___closed__8; @@ -1094,13 +1075,13 @@ static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabDIte___cl LEAN_EXPORT lean_object* l_Lean_unresolveNameGlobal___at_Lean_PrettyPrinter_Delaborator_delabConst___spec__1___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_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLam___lambda__1___closed__3; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppFn(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabOfNat___closed__4; uint8_t l_Lean_isStructure(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabProj___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_mkArray1___rarg(lean_object*); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__18(lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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___regBuiltin_Lean_PrettyPrinter_Delaborator_coeDelaborator___closed__1; lean_object* l_Lean_Parser_termParser_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_declSigWithId_formatter___closed__7; @@ -1115,10 +1096,13 @@ static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabDIte___cl LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabForall___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_withOverApp(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_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabDIte___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__7___closed__2; lean_object* l_Lean_PrettyPrinter_Formatter_andthen_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__2___closed__3; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___closed__3; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__4; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDIte___lambda__1___boxed(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*); @@ -1126,8 +1110,8 @@ static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabSort___cl static lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__3___closed__10; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLit___closed__3; 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_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__15(lean_object*, lean_object*, size_t, size_t, lean_object*); extern lean_object* l_Lean_Parser_maxPrec; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__5(lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDivRatCore___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_reifyName(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabProjectionApp_withoutParentProjections___spec__1(lean_object*); @@ -1138,32 +1122,29 @@ static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabLit___clo static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabSort___closed__2; lean_object* l_Array_ofSubarray___rarg(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabProjectionApp___closed__4; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__13(lean_object*, lean_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_delabSigmaCore___lambda__1___closed__1; uint8_t l_Lean_LocalDecl_binderInfo(lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabSort___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabOfScientific___closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_useAppExplicit___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_useAppExplicit___lambda__2___closed__1; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__19___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__20(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabMData___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabNeg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabLetFun___closed__1; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__41___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_erase_macro_scopes(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabOfScientific___lambda__1___closed__3; lean_object* l_Array_back___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__11(lean_object*, lean_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_delabMData___spec__11(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_delabAppImplicitCore___closed__2; LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabMData___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_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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* l_Lean_getPPNatLit___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__14(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda__11___closed__2; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__6; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabNeg___closed__2; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabApp___closed__1; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabSigma___closed__2; @@ -1179,9 +1160,7 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__1_ static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabLam___closed__5; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabNameMkStr___closed__2; -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___closed__1; -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_PrettyPrinter_Delaborator_delabCond___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDIte_delabBranch___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_delabAppExplicitCore___lambda__4___closed__1; @@ -1193,7 +1172,6 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___closed__4; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabForall___closed__4; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabMData___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppCore___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__11(lean_object*, size_t, size_t, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabLit___closed__5; LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabApp_delabAppFn(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1217,15 +1195,14 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabLam___lambda__2__ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__7; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__1; +LEAN_EXPORT uint8_t l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_shouldGroupWithNext(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabProjectionApp_withoutParentProjections___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabDoElems___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_withMDatasOptions___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_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__11(lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabForall___closed__2; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabNegIntCore___closed__6; lean_object* l_Lean_getPPLetVarTypes___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_PrettyPrinter_Delaborator_annotateCurPos(lean_object*, lean_object*, 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_Delaborator_useAppExplicit___closed__1; @@ -1251,9 +1228,10 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambd LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___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* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__8___closed__2; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__26(lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__12(lean_object*, lean_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_delabNameMkStr___lambda__1___closed__2; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_shouldGroupWithNext___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDo___lambda__1___closed__4; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDivRatCore___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1261,30 +1239,28 @@ lean_object* lean_string_append(lean_object*, lean_object*); lean_object* l_Lean_KeyedDeclsAttribute_getValues___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabProjectionApp_withoutParentProjections(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFnArgs___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__6(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabForall___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withProj___at_Lean_PrettyPrinter_Delaborator_delabProj___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__12(lean_object*, lean_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_delabLetFun___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConst___closed__7; LEAN_EXPORT lean_object* l_Lean_resolveGlobalName___at_Lean_PrettyPrinter_Delaborator_delabConst___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabSort___closed__1; -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabForall___lambda__2___closed__1; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabChar___closed__1; lean_object* l_Lean_PrettyPrinter_Delaborator_withBindingBodyUnusedName___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_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__3; lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__16___closed__2; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isMData(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_withMDataOptions(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_withTypeAscription(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_delabLam___lambda__1___closed__5; +static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___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_PrettyPrinter_Delaborator_delabDoElems___lambda__8___closed__5; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabProjectionApp___closed__2; @@ -1292,12 +1268,10 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgS LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabSigmaCore___lambda__3(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_delabApp_delabAppFn___closed__4; lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabSigma___closed__3; uint8_t lean_nat_dec_le(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__4___closed__3; static lean_object* l_Lean_PrettyPrinter_Delaborator_declSigWithId_parenthesizer___closed__2; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__18___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabOfNat___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6(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); @@ -1359,12 +1333,10 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda__6___ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabLam(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda__8(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_delabMData___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_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__16(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__3(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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* l_Lean_Meta_isConstructorApp_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLam___closed__1; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__3(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__4; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_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*); @@ -1386,10 +1358,9 @@ lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_Delaborator_whenPPOption(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_expr_instantiate1(lean_object*, lean_object*); lean_object* l_Lean_getPPAnalysisNamedArg___boxed(lean_object*); +static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__7___closed__4; lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_instantiateForallAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__4(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withProj___at_Lean_PrettyPrinter_Delaborator_delabProj___spec__1___closed__2; -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__14(lean_object*, 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_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_evalSyntaxConstantUnsafe___closed__1; LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__7___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__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*); @@ -1412,6 +1383,7 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabCond___lambda__1___clo LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppCore(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_PrettyPrinter_Delaborator_delabConst___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_SubExpr_withProj___at_Lean_PrettyPrinter_Delaborator_delabProj___spec__1___closed__1; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__18(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_unresolveNameGlobal_unresolveNameCore___at_Lean_PrettyPrinter_Delaborator_delabConst___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_AppMatchState_rhss___default; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConst___closed__12; @@ -1424,6 +1396,7 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_PrettyPrinter_Delabo LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_shouldGroupWithNext___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabLam(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabDIte___closed__5; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDIte___lambda__1___closed__4; @@ -39498,295 +39471,175 @@ lean_dec(x_2); return x_4; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__1(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__1() { _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; size_t x_9; size_t x_10; lean_object* x_11; -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 = 1; -x_10 = lean_usize_add(x_3, x_9); -x_11 = lean_array_uset(x_8, x_3, x_6); -x_3 = x_10; -x_4 = x_11; -goto _start; -} +lean_object* x_1; +x_1 = l_Lean_pp_piBinderTypes; +return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__2(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy(lean_object* x_1, 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_5; -x_5 = lean_usize_dec_lt(x_3, x_2); -if (x_5 == 0) +uint8_t x_8; +x_8 = !lean_is_exclusive(x_5); +if (x_8 == 0) { -return x_4; +lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; lean_object* x_13; +x_9 = lean_ctor_get(x_5, 2); +x_10 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__1; +x_11 = 1; +x_12 = l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(x_9, x_10, x_11); +lean_ctor_set(x_5, 2, x_12); +x_13 = l_Lean_PrettyPrinter_Delaborator_delab(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +return x_13; } else { -lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; -x_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 = 1; -x_10 = lean_usize_add(x_3, x_9); -x_11 = lean_array_uset(x_8, x_3, x_6); -x_3 = x_10; -x_4 = x_11; -goto _start; +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; lean_object* x_26; uint8_t x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_14 = lean_ctor_get(x_5, 0); +x_15 = lean_ctor_get(x_5, 1); +x_16 = lean_ctor_get(x_5, 2); +x_17 = lean_ctor_get(x_5, 3); +x_18 = lean_ctor_get(x_5, 4); +x_19 = lean_ctor_get(x_5, 5); +x_20 = lean_ctor_get(x_5, 6); +x_21 = lean_ctor_get(x_5, 7); +x_22 = lean_ctor_get(x_5, 8); +x_23 = lean_ctor_get(x_5, 9); +x_24 = lean_ctor_get(x_5, 10); +x_25 = lean_ctor_get_uint8(x_5, sizeof(void*)*11); +lean_inc(x_24); +lean_inc(x_23); +lean_inc(x_22); +lean_inc(x_21); +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_dec(x_5); +x_26 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__1; +x_27 = 1; +x_28 = l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(x_16, x_26, x_27); +x_29 = lean_alloc_ctor(0, 11, 1); +lean_ctor_set(x_29, 0, x_14); +lean_ctor_set(x_29, 1, x_15); +lean_ctor_set(x_29, 2, x_28); +lean_ctor_set(x_29, 3, x_17); +lean_ctor_set(x_29, 4, x_18); +lean_ctor_set(x_29, 5, x_19); +lean_ctor_set(x_29, 6, x_20); +lean_ctor_set(x_29, 7, x_21); +lean_ctor_set(x_29, 8, x_22); +lean_ctor_set(x_29, 9, x_23); +lean_ctor_set(x_29, 10, x_24); +lean_ctor_set_uint8(x_29, sizeof(void*)*11, x_25); +x_30 = l_Lean_PrettyPrinter_Delaborator_delab(x_1, x_2, x_3, x_4, x_29, x_6, x_7); +return x_30; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__3(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +LEAN_EXPORT uint8_t l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_shouldGroupWithNext(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _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; size_t x_9; size_t x_10; lean_object* x_11; -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 = 1; -x_10 = lean_usize_add(x_3, x_9); -x_11 = lean_array_uset(x_8, x_3, x_6); -x_3 = x_10; -x_4 = x_11; -goto _start; -} -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__4(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { -_start: +uint8_t x_4; +x_4 = l_Lean_Expr_isForall(x_3); +if (x_4 == 0) { uint8_t x_5; -x_5 = lean_usize_dec_lt(x_3, x_2); -if (x_5 == 0) -{ -return x_4; +x_5 = 0; +return x_5; } else { -lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; -x_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 = 1; -x_10 = lean_usize_add(x_3, x_9); -x_11 = lean_array_uset(x_8, x_3, x_6); -x_3 = x_10; -x_4 = x_11; -goto _start; -} -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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; uint8_t x_7; +x_6 = l_Lean_Expr_bindingName_x21(x_3); +x_7 = l_Lean_Name_hasMacroScopes(x_6); +if (x_7 == 0) { -lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; -x_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 = 1; -x_10 = lean_usize_add(x_3, x_9); -x_11 = lean_array_uset(x_8, x_3, x_6); -x_3 = x_10; -x_4 = x_11; -goto _start; -} -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__6(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { -_start: +uint8_t x_8; +x_8 = l_Lean_NameSet_contains(x_1, x_6); +lean_dec(x_6); +if (x_8 == 0) { -uint8_t x_5; -x_5 = lean_usize_dec_lt(x_3, x_2); -if (x_5 == 0) +uint8_t x_9; uint8_t x_10; uint8_t x_11; +x_9 = l_Lean_Expr_binderInfo(x_2); +x_10 = l_Lean_Expr_binderInfo(x_3); +x_11 = l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_399_(x_9, x_10); +if (x_11 == 0) { -return x_4; +uint8_t x_12; +x_12 = 0; +return x_12; } else { -lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; -x_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 = 1; -x_10 = lean_usize_add(x_3, x_9); -x_11 = lean_array_uset(x_8, x_3, x_6); -x_3 = x_10; -x_4 = x_11; -goto _start; -} -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__7(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) +lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_13 = l_Lean_Expr_bindingDomain_x21(x_2); +x_14 = l_Lean_Expr_bindingDomain_x21(x_3); +x_15 = lean_expr_eqv(x_13, x_14); +lean_dec(x_14); +lean_dec(x_13); +if (x_15 == 0) { -return x_4; +uint8_t x_16; +x_16 = 0; +return x_16; } else { -lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; -x_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 = 1; -x_10 = lean_usize_add(x_3, x_9); -x_11 = lean_array_uset(x_8, x_3, x_6); -x_3 = x_10; -x_4 = x_11; -goto _start; -} -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__8(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) +uint8_t x_17; uint8_t x_18; +x_17 = 3; +x_18 = l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_399_(x_10, x_17); +if (x_18 == 0) { -return x_4; +uint8_t x_19; +x_19 = 1; +return x_19; } else { -lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; -x_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 = 1; -x_10 = lean_usize_add(x_3, x_9); -x_11 = lean_array_uset(x_8, x_3, x_6); -x_3 = x_10; -x_4 = x_11; -goto _start; +uint8_t x_20; +x_20 = 0; +return x_20; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__9(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; size_t x_9; size_t x_10; lean_object* x_11; -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 = 1; -x_10 = lean_usize_add(x_3, x_9); -x_11 = lean_array_uset(x_8, x_3, x_6); -x_3 = x_10; -x_4 = x_11; -goto _start; -} -} +uint8_t x_21; +x_21 = 0; +return x_21; } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__10(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; size_t x_9; size_t x_10; lean_object* x_11; -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 = 1; -x_10 = lean_usize_add(x_3, x_9); -x_11 = lean_array_uset(x_8, x_3, x_6); -x_3 = x_10; -x_4 = x_11; -goto _start; -} -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__11(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; +uint8_t x_22; +lean_dec(x_6); +x_22 = 0; +return x_22; } -else -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; -x_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 = 1; -x_10 = lean_usize_add(x_3, x_9); -x_11 = lean_array_uset(x_8, x_3, x_6); -x_3 = x_10; -x_4 = x_11; -goto _start; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__12(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_shouldGroupWithNext___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _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; size_t x_9; size_t x_10; lean_object* x_11; -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 = 1; -x_10 = lean_usize_add(x_3, x_9); -x_11 = lean_array_uset(x_8, x_3, x_6); -x_3 = x_10; -x_4 = x_11; -goto _start; -} +uint8_t x_4; lean_object* x_5; +x_4 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_shouldGroupWithNext(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_5 = lean_box(x_4); +return x_5; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__13(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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__1(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { uint8_t x_5; @@ -39810,7 +39663,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__14(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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__2(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { uint8_t x_5; @@ -39834,31 +39687,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__15(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_lt(x_4, x_3); -if (x_6 == 0) -{ -return x_5; -} -else -{ -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_7 = lean_array_uget(x_5, x_4); -x_8 = lean_unsigned_to_nat(0u); -x_9 = lean_array_uset(x_5, x_4, x_8); -x_10 = 1; -x_11 = lean_usize_add(x_4, x_10); -x_12 = lean_array_uset(x_9, x_4, x_7); -x_4 = x_11; -x_5 = x_12; -goto _start; -} -} -} -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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: { uint8_t x_11; @@ -39914,7 +39743,7 @@ return x_27; } } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingDomain___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__16(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingDomain___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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; lean_object* x_13; lean_object* x_14; @@ -39927,174 +39756,459 @@ lean_dec(x_9); x_12 = l_Lean_Expr_bindingDomain_x21(x_10); lean_dec(x_10); x_13 = lean_unsigned_to_nat(0u); -x_14 = l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__17(x_12, x_13, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_11); +x_14 = l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__4(x_12, x_13, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_11); return x_14; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__18(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_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_lt(x_4, x_3); -if (x_6 == 0) -{ -return x_5; +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; lean_object* x_21; +x_9 = lean_ctor_get(x_6, 5); +lean_inc(x_9); +lean_dec(x_6); +x_10 = 0; +x_11 = l_Lean_SourceInfo_fromRef(x_9, x_10); +x_12 = l_Lean_PrettyPrinter_Delaborator_delabLam___lambda__3___closed__4; +lean_inc(x_11); +x_13 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_13, 0, x_11); +lean_ctor_set(x_13, 1, x_12); +x_14 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; +x_15 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; +lean_inc(x_11); +x_16 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_16, 0, x_11); +lean_ctor_set(x_16, 1, x_14); +lean_ctor_set(x_16, 2, x_15); +x_17 = l_Lean_PrettyPrinter_Delaborator_delabLam___lambda__3___closed__5; +lean_inc(x_11); +x_18 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_18, 0, x_11); +lean_ctor_set(x_18, 1, x_17); +x_19 = l_Lean_PrettyPrinter_Delaborator_delabLam___lambda__3___closed__3; +x_20 = l_Lean_Syntax_node4(x_11, x_19, x_13, x_16, x_1, 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_8); +return x_21; } -else -{ -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_7 = lean_array_uget(x_5, x_4); -x_8 = lean_unsigned_to_nat(0u); -x_9 = lean_array_uset(x_5, x_4, x_8); -x_10 = 1; -x_11 = lean_usize_add(x_4, x_10); -x_12 = lean_array_uset(x_9, x_4, x_7); -x_4 = x_11; -x_5 = x_12; -goto _start; } +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Command", 7); +return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__19(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_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__2() { _start: { -uint8_t x_6; -x_6 = lean_usize_dec_lt(x_4, x_3); -if (x_6 == 0) +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("declSig", 7); +return x_1; +} +} +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_PrettyPrinter_Delaborator_withTypeAscription___closed__2; +x_2 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__3; +x_3 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__1; +x_4 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__2; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } -else -{ -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_7 = lean_array_uget(x_5, x_4); -x_8 = lean_unsigned_to_nat(0u); -x_9 = lean_array_uset(x_5, x_4, x_8); -x_10 = 1; -x_11 = lean_usize_add(x_4, x_10); -x_12 = lean_array_uset(x_9, x_4, x_7); -x_4 = x_11; -x_5 = x_12; -goto _start; } +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("ident", 5); +return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__20(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_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5() { _start: { -uint8_t x_6; -x_6 = lean_usize_dec_lt(x_4, x_3); -if (x_6 == 0) -{ -return x_5; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__4; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } -else -{ -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_7 = lean_array_uget(x_5, x_4); -x_8 = lean_unsigned_to_nat(0u); -x_9 = lean_array_uset(x_5, x_4, x_8); -x_10 = 1; -x_11 = lean_usize_add(x_4, x_10); -x_12 = lean_array_uset(x_9, x_4, x_7); -x_4 = x_11; -x_5 = x_12; -goto _start; } +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("bracketedBinder", 15); +return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__21(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_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__7() { _start: { -uint8_t x_6; -x_6 = lean_usize_dec_lt(x_4, x_3); -if (x_6 == 0) -{ +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_Delaborator_withTypeAscription___closed__2; +x_2 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__3; +x_3 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__4; +x_4 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__6; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } -else -{ -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_7 = lean_array_uget(x_5, x_4); -x_8 = lean_unsigned_to_nat(0u); -x_9 = lean_array_uset(x_5, x_4, x_8); -x_10 = 1; -x_11 = lean_usize_add(x_4, x_10); -x_12 = lean_array_uset(x_9, x_4, x_7); -x_4 = x_11; -x_5 = x_12; -goto _start; } +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__7; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__23(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9() { _start: { -lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_9 = lean_ctor_get(x_6, 5); -x_10 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_4, x_5, x_6, x_7, x_8); -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_inc(x_9); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_9); -lean_ctor_set(x_13, 1, x_12); -lean_ctor_set_tag(x_10, 1); -lean_ctor_set(x_10, 0, x_13); -return x_10; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__4___closed__2; +x_2 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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; } -else +} +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__10() { +_start: { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_14 = lean_ctor_get(x_10, 0); -x_15 = lean_ctor_get(x_10, 1); -lean_inc(x_15); -lean_inc(x_14); -lean_dec(x_10); -lean_inc(x_9); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_9); -lean_ctor_set(x_16, 1, x_14); -x_17 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set(x_17, 1, x_15); -return x_17; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5; +x_2 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy), 7, 0); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_ofExcept___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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) { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__12() { _start: { -if (lean_obj_tag(x_1) == 0) +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__1___boxed), 8, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__13() { +_start: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_9 = lean_ctor_get(x_1, 0); -lean_inc(x_9); +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11; +x_2 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__12; +x_3 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__3___rarg), 9, 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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, 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_74; +x_11 = l_Lean_PrettyPrinter_Delaborator_SubExpr_getExpr___at_Lean_PrettyPrinter_Delaborator_getExprKind___spec__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_74 = l_Lean_Expr_isForall(x_12); +if (x_74 == 0) +{ +lean_object* x_75; +lean_dec(x_12); lean_dec(x_1); -x_10 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_10, 0, x_9); -x_11 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_11, 0, x_10); -x_12 = l_Lean_throwError___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__23(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_12; +x_75 = lean_box(0); +x_14 = x_75; +goto block_73; } else { -lean_object* x_13; lean_object* x_14; -x_13 = lean_ctor_get(x_1, 0); -lean_inc(x_13); +uint8_t x_76; uint8_t x_77; +x_76 = l_Lean_Expr_binderInfo(x_12); +x_77 = l_Lean_BinderInfo_isInstImplicit(x_76); +if (x_77 == 0) +{ +lean_object* x_78; uint8_t x_79; +x_78 = l_Lean_Expr_bindingName_x21(x_12); +lean_dec(x_12); +x_79 = l_Lean_Name_hasMacroScopes(x_78); +if (x_79 == 0) +{ +uint8_t x_80; +x_80 = l_Lean_NameSet_contains(x_1, x_78); +lean_dec(x_78); +if (x_80 == 0) +{ +lean_object* x_81; lean_object* x_82; +x_81 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; +x_82 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux(x_1, x_2, x_3, x_81, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_82; +} +else +{ +lean_object* x_83; lean_dec(x_1); -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; +x_83 = lean_box(0); +x_14 = x_83; +goto block_73; +} +} +else +{ +lean_object* x_84; +lean_dec(x_78); +lean_dec(x_1); +x_84 = lean_box(0); +x_14 = x_84; +goto block_73; +} +} +else +{ +lean_object* x_85; uint8_t x_86; +x_85 = l_Lean_Expr_bindingName_x21(x_12); +lean_dec(x_12); +x_86 = l_Lean_Name_hasMacroScopes(x_85); +if (x_86 == 0) +{ +uint8_t x_87; +x_87 = l_Lean_NameSet_contains(x_1, x_85); +lean_dec(x_85); +if (x_87 == 0) +{ +lean_object* x_88; lean_object* x_89; +x_88 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; +x_89 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux(x_1, x_2, x_3, x_88, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_89; +} +else +{ +lean_object* x_90; +lean_dec(x_1); +x_90 = lean_box(0); +x_14 = x_90; +goto block_73; +} +} +else +{ +lean_object* x_91; lean_object* x_92; +x_91 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__13; +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_92 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingDomain___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__3(x_91, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +if (lean_obj_tag(x_92) == 0) +{ +lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; +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_array_push(x_3, x_93); +x_96 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams), 10, 3); +lean_closure_set(x_96, 0, x_1); +lean_closure_set(x_96, 1, x_2); +lean_closure_set(x_96, 2, x_95); +x_97 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(x_85, x_96, x_4, x_5, x_6, x_7, x_8, x_9, x_94); +return x_97; +} +else +{ +uint8_t x_98; +lean_dec(x_85); +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_98 = !lean_is_exclusive(x_92); +if (x_98 == 0) +{ +return x_92; +} +else +{ +lean_object* x_99; lean_object* x_100; lean_object* x_101; +x_99 = lean_ctor_get(x_92, 0); +x_100 = lean_ctor_get(x_92, 1); +lean_inc(x_100); +lean_inc(x_99); +lean_dec(x_92); +x_101 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_101, 0, x_99); +lean_ctor_set(x_101, 1, x_100); +return x_101; +} +} +} +} +} +block_73: +{ +lean_object* x_15; +lean_dec(x_14); +lean_inc(x_8); +x_15 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy(x_4, x_5, x_6, x_7, x_8, x_9, x_13); +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; uint8_t 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; size_t x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_17 = lean_ctor_get(x_15, 0); +x_18 = lean_ctor_get(x_8, 5); +lean_inc(x_18); +lean_dec(x_8); +x_19 = 0; +x_20 = l_Lean_SourceInfo_fromRef(x_18, x_19); +x_21 = lean_array_get_size(x_3); +x_22 = lean_usize_of_nat(x_21); +lean_dec(x_21); +x_23 = 0; +x_24 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; +x_25 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__1(x_24, x_22, x_23, x_3); +x_26 = lean_array_get_size(x_25); +x_27 = lean_usize_of_nat(x_26); +lean_dec(x_26); +x_28 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__10; +x_29 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__2(x_28, x_27, x_23, x_25); +x_30 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; +x_31 = l_Array_append___rarg(x_30, x_29); +x_32 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; +lean_inc(x_20); +x_33 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_33, 0, x_20); +lean_ctor_set(x_33, 1, x_32); +lean_ctor_set(x_33, 2, x_31); +x_34 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__8; +lean_inc(x_20); +x_35 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_35, 0, x_20); +lean_ctor_set(x_35, 1, x_34); +x_36 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__9; +lean_inc(x_20); +x_37 = l_Lean_Syntax_node2(x_20, x_36, x_35, x_17); +x_38 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__3; +lean_inc(x_20); +x_39 = l_Lean_Syntax_node2(x_20, x_38, x_33, x_37); +x_40 = l_Lean_PrettyPrinter_Delaborator_declSigWithId_formatter___closed__2; +x_41 = l_Lean_Syntax_node2(x_20, x_40, x_2, x_39); +lean_ctor_set(x_15, 0, x_41); +return x_15; +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; lean_object* x_46; lean_object* x_47; size_t x_48; size_t x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; size_t x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_42 = lean_ctor_get(x_15, 0); +x_43 = lean_ctor_get(x_15, 1); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_15); +x_44 = lean_ctor_get(x_8, 5); +lean_inc(x_44); +lean_dec(x_8); +x_45 = 0; +x_46 = l_Lean_SourceInfo_fromRef(x_44, x_45); +x_47 = lean_array_get_size(x_3); +x_48 = lean_usize_of_nat(x_47); +lean_dec(x_47); +x_49 = 0; +x_50 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; +x_51 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__1(x_50, x_48, x_49, x_3); +x_52 = lean_array_get_size(x_51); +x_53 = lean_usize_of_nat(x_52); +lean_dec(x_52); +x_54 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__10; +x_55 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__2(x_54, x_53, x_49, x_51); +x_56 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; +x_57 = l_Array_append___rarg(x_56, x_55); +x_58 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; +lean_inc(x_46); +x_59 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_59, 0, x_46); +lean_ctor_set(x_59, 1, x_58); +lean_ctor_set(x_59, 2, x_57); +x_60 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__8; +lean_inc(x_46); +x_61 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_61, 0, x_46); +lean_ctor_set(x_61, 1, x_60); +x_62 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__9; +lean_inc(x_46); +x_63 = l_Lean_Syntax_node2(x_46, x_62, x_61, x_42); +x_64 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__3; +lean_inc(x_46); +x_65 = l_Lean_Syntax_node2(x_46, x_64, x_59, x_63); +x_66 = l_Lean_PrettyPrinter_Delaborator_declSigWithId_formatter___closed__2; +x_67 = l_Lean_Syntax_node2(x_46, x_66, x_2, x_65); +x_68 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_68, 0, x_67); +lean_ctor_set(x_68, 1, x_43); +return x_68; +} +} +else +{ +uint8_t x_69; +lean_dec(x_8); +lean_dec(x_3); +lean_dec(x_2); +x_69 = !lean_is_exclusive(x_15); +if (x_69 == 0) +{ +return x_15; +} +else +{ +lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_70 = lean_ctor_get(x_15, 0); +x_71 = lean_ctor_get(x_15, 1); +lean_inc(x_71); +lean_inc(x_70); +lean_dec(x_15); +x_72 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_72, 0, x_70); +lean_ctor_set(x_72, 1, x_71); +return x_72; +} } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__24(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__1(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; @@ -40118,7 +40232,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__25(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__2(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; @@ -40142,7 +40256,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__26(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__3(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; @@ -40166,7 +40280,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__27(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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; @@ -40190,7 +40304,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__28(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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; @@ -40214,7 +40328,74 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__29(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_Lean_throwError___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_9 = lean_ctor_get(x_6, 5); +x_10 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_4, x_5, x_6, x_7, x_8); +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_inc(x_9); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_9); +lean_ctor_set(x_13, 1, x_12); +lean_ctor_set_tag(x_10, 1); +lean_ctor_set(x_10, 0, x_13); +return x_10; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_14 = lean_ctor_get(x_10, 0); +x_15 = lean_ctor_get(x_10, 1); +lean_inc(x_15); +lean_inc(x_14); +lean_dec(x_10); +lean_inc(x_9); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_9); +lean_ctor_set(x_16, 1, x_14); +x_17 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_15); +return x_17; +} +} +} +LEAN_EXPORT lean_object* l_Lean_ofExcept___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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; lean_object* x_11; lean_object* x_12; +x_9 = lean_ctor_get(x_1, 0); +lean_inc(x_9); +lean_dec(x_1); +x_10 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_10, 0, x_9); +x_11 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_11, 0, x_10); +x_12 = l_Lean_throwError___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__7(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +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); +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; +} +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__8(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; @@ -40238,7 +40419,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__30(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__9(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; @@ -40262,7 +40443,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__31(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__10(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; @@ -40286,7 +40467,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__32(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__11(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; @@ -40310,7 +40491,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__33(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__12(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { _start: { uint8_t x_6; @@ -40334,7 +40515,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__34(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__13(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; @@ -40358,255 +40539,127 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__35(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; size_t x_9; size_t x_10; lean_object* x_11; -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 = 1; -x_10 = lean_usize_add(x_3, x_9); -x_11 = lean_array_uset(x_8, x_3, x_6); -x_3 = x_10; -x_4 = x_11; -goto _start; -} -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__36(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; size_t x_9; size_t x_10; lean_object* x_11; -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 = 1; -x_10 = lean_usize_add(x_3, x_9); -x_11 = lean_array_uset(x_8, x_3, x_6); -x_3 = x_10; -x_4 = x_11; -goto _start; -} -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__37(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; size_t x_9; size_t x_10; lean_object* x_11; -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 = 1; -x_10 = lean_usize_add(x_3, x_9); -x_11 = lean_array_uset(x_8, x_3, x_6); -x_3 = x_10; -x_4 = x_11; -goto _start; -} -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__38(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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__14(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { _start: { -uint8_t x_5; -x_5 = lean_usize_dec_lt(x_3, x_2); -if (x_5 == 0) +uint8_t x_6; +x_6 = lean_usize_dec_lt(x_4, x_3); +if (x_6 == 0) { -return x_4; +return x_5; } else { -lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; -x_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 = 1; -x_10 = lean_usize_add(x_3, x_9); -x_11 = lean_array_uset(x_8, x_3, x_6); -x_3 = x_10; +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_7 = lean_array_uget(x_5, x_4); +x_8 = lean_unsigned_to_nat(0u); +x_9 = lean_array_uset(x_5, x_4, x_8); +x_10 = 1; +x_11 = lean_usize_add(x_4, x_10); +x_12 = lean_array_uset(x_9, x_4, x_7); x_4 = x_11; +x_5 = x_12; goto _start; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__39(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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__15(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { _start: { -uint8_t x_5; -x_5 = lean_usize_dec_lt(x_3, x_2); -if (x_5 == 0) +uint8_t x_6; +x_6 = lean_usize_dec_lt(x_4, x_3); +if (x_6 == 0) { -return x_4; +return x_5; } else { -lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; -x_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 = 1; -x_10 = lean_usize_add(x_3, x_9); -x_11 = lean_array_uset(x_8, x_3, x_6); -x_3 = x_10; +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_7 = lean_array_uget(x_5, x_4); +x_8 = lean_unsigned_to_nat(0u); +x_9 = lean_array_uset(x_5, x_4, x_8); +x_10 = 1; +x_11 = lean_usize_add(x_4, x_10); +x_12 = lean_array_uset(x_9, x_4, x_7); x_4 = x_11; +x_5 = x_12; goto _start; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__40(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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__16(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { _start: { -uint8_t x_5; -x_5 = lean_usize_dec_lt(x_3, x_2); -if (x_5 == 0) +uint8_t x_6; +x_6 = lean_usize_dec_lt(x_4, x_3); +if (x_6 == 0) { -return x_4; +return x_5; } else { -lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; -x_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 = 1; -x_10 = lean_usize_add(x_3, x_9); -x_11 = lean_array_uset(x_8, x_3, x_6); -x_3 = x_10; +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_7 = lean_array_uget(x_5, x_4); +x_8 = lean_unsigned_to_nat(0u); +x_9 = lean_array_uset(x_5, x_4, x_8); +x_10 = 1; +x_11 = lean_usize_add(x_4, x_10); +x_12 = lean_array_uset(x_9, x_4, x_7); x_4 = x_11; +x_5 = x_12; goto _start; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__41(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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__17(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { _start: { -uint8_t x_5; -x_5 = lean_usize_dec_lt(x_3, x_2); -if (x_5 == 0) +uint8_t x_6; +x_6 = lean_usize_dec_lt(x_4, x_3); +if (x_6 == 0) { -return x_4; +return x_5; } else { -lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; -x_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 = 1; -x_10 = lean_usize_add(x_3, x_9); -x_11 = lean_array_uset(x_8, x_3, x_6); -x_3 = x_10; +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_7 = lean_array_uget(x_5, x_4); +x_8 = lean_unsigned_to_nat(0u); +x_9 = lean_array_uset(x_5, x_4, x_8); +x_10 = 1; +x_11 = lean_usize_add(x_4, x_10); +x_12 = lean_array_uset(x_9, x_4, x_7); x_4 = x_11; +x_5 = x_12; goto _start; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__42(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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__18(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { _start: { -uint8_t x_5; -x_5 = lean_usize_dec_lt(x_3, x_2); -if (x_5 == 0) +uint8_t x_6; +x_6 = lean_usize_dec_lt(x_4, x_3); +if (x_6 == 0) { -return x_4; +return x_5; } else { -lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; -x_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 = 1; -x_10 = lean_usize_add(x_3, x_9); -x_11 = lean_array_uset(x_8, x_3, x_6); -x_3 = x_10; +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_7 = lean_array_uget(x_5, x_4); +x_8 = lean_unsigned_to_nat(0u); +x_9 = lean_array_uset(x_5, x_4, x_8); +x_10 = 1; +x_11 = lean_usize_add(x_4, x_10); +x_12 = lean_array_uset(x_9, x_4, x_7); x_4 = x_11; +x_5 = x_12; goto _start; } } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__1___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_pp_piBinderTypes; -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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; uint8_t x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_23; lean_object* x_24; -x_8 = lean_ctor_get(x_5, 0); -lean_inc(x_8); -x_9 = lean_ctor_get(x_5, 1); -lean_inc(x_9); -x_10 = lean_ctor_get(x_5, 2); -lean_inc(x_10); -x_11 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__1___closed__1; -x_12 = 1; -x_13 = l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(x_10, x_11, x_12); -x_14 = lean_ctor_get(x_5, 3); -lean_inc(x_14); -x_15 = lean_ctor_get(x_5, 4); -lean_inc(x_15); -x_16 = lean_ctor_get(x_5, 5); -lean_inc(x_16); -x_17 = lean_ctor_get(x_5, 6); -lean_inc(x_17); -x_18 = lean_ctor_get(x_5, 7); -lean_inc(x_18); -x_19 = lean_ctor_get(x_5, 8); -lean_inc(x_19); -x_20 = lean_ctor_get(x_5, 9); -lean_inc(x_20); -x_21 = lean_ctor_get(x_5, 10); -lean_inc(x_21); -x_22 = lean_ctor_get_uint8(x_5, sizeof(void*)*11); -lean_dec(x_5); -x_23 = lean_alloc_ctor(0, 11, 1); -lean_ctor_set(x_23, 0, x_8); -lean_ctor_set(x_23, 1, x_9); -lean_ctor_set(x_23, 2, x_13); -lean_ctor_set(x_23, 3, x_14); -lean_ctor_set(x_23, 4, x_15); -lean_ctor_set(x_23, 5, x_16); -lean_ctor_set(x_23, 6, x_17); -lean_ctor_set(x_23, 7, x_18); -lean_ctor_set(x_23, 8, x_19); -lean_ctor_set(x_23, 9, x_20); -lean_ctor_set(x_23, 10, x_21); -lean_ctor_set_uint8(x_23, sizeof(void*)*11, x_22); -x_24 = l_Lean_PrettyPrinter_Delaborator_delab(x_1, x_2, x_3, x_4, x_23, x_6, x_7); -return x_24; -} -} -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_12; uint8_t x_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; 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; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; @@ -40628,7 +40681,7 @@ x_19 = lean_array_get_size(x_2); x_20 = lean_usize_of_nat(x_19); lean_dec(x_19); x_21 = 0; -x_22 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__15(x_3, x_18, x_20, x_21, x_2); +x_22 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__1(x_3, x_18, x_20, x_21, x_2); lean_dec(x_18); x_23 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; x_24 = l_Array_append___rarg(x_23, x_22); @@ -40663,7 +40716,7 @@ lean_ctor_set(x_35, 1, x_11); return x_35; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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; uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; 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; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; @@ -40685,7 +40738,7 @@ x_19 = lean_array_get_size(x_2); x_20 = lean_usize_of_nat(x_19); lean_dec(x_19); x_21 = 0; -x_22 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__18(x_3, x_18, x_20, x_21, x_2); +x_22 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__2(x_3, x_18, x_20, x_21, x_2); lean_dec(x_18); x_23 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; x_24 = l_Array_append___rarg(x_23, x_22); @@ -40720,7 +40773,7 @@ lean_ctor_set(x_35, 1, x_11); return x_35; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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) { _start: { lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; 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; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; @@ -40742,7 +40795,7 @@ x_19 = lean_array_get_size(x_2); x_20 = lean_usize_of_nat(x_19); lean_dec(x_19); x_21 = 0; -x_22 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__19(x_3, x_18, x_20, x_21, x_2); +x_22 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__3(x_3, x_18, x_20, x_21, x_2); lean_dec(x_18); x_23 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; x_24 = l_Array_append___rarg(x_23, x_22); @@ -40777,7 +40830,7 @@ lean_ctor_set(x_35, 1, x_11); return x_35; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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; uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; 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; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; @@ -40799,7 +40852,7 @@ x_19 = lean_array_get_size(x_2); x_20 = lean_usize_of_nat(x_19); lean_dec(x_19); x_21 = 0; -x_22 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__20(x_3, x_18, x_20, x_21, x_2); +x_22 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__4(x_3, x_18, x_20, x_21, x_2); lean_dec(x_18); x_23 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; x_24 = l_Array_append___rarg(x_23, x_22); @@ -40834,7 +40887,7 @@ lean_ctor_set(x_35, 1, x_11); return x_35; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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; uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; 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; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; @@ -40856,7 +40909,7 @@ x_19 = lean_array_get_size(x_2); x_20 = lean_usize_of_nat(x_19); lean_dec(x_19); x_21 = 0; -x_22 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__21(x_3, x_18, x_20, x_21, x_2); +x_22 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__5(x_3, x_18, x_20, x_21, x_2); lean_dec(x_18); x_23 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; x_24 = l_Array_append___rarg(x_23, x_22); @@ -40891,7 +40944,7 @@ lean_ctor_set(x_35, 1, x_11); return x_35; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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; uint8_t x_9; @@ -40925,7 +40978,17 @@ return x_15; } } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__8___closed__1() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__7___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11; +x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__1), 8, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__7___closed__2() { _start: { lean_object* x_1; @@ -40933,19 +40996,19 @@ x_1 = lean_mk_string_from_bytes("binderTactic", 12); return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__8___closed__2() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__7___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_PrettyPrinter_Delaborator_withTypeAscription___closed__2; x_2 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__3; x_3 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__4; -x_4 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__8___closed__1; +x_4 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__7___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__8___closed__3() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__7___closed__4() { _start: { lean_object* x_1; @@ -40953,227 +41016,225 @@ x_1 = lean_mk_string_from_bytes("by", 2); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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) { _start: { -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_11, 2); -lean_inc(x_14); -x_15 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_evalSyntaxConstantUnsafe(x_6, x_14, x_1); -lean_dec(x_14); -x_16 = l_Lean_ofExcept___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__22(x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_16) == 0) +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_10, 2); +lean_inc(x_13); +x_14 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_evalSyntaxConstantUnsafe(x_5, x_13, x_1); +lean_dec(x_13); +x_15 = l_Lean_ofExcept___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__6(x_14, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_17 = lean_ctor_get(x_16, 0); +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); -lean_inc(x_18); -lean_dec(x_16); -x_19 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__1), 8, 1); -lean_closure_set(x_19, 0, x_2); -lean_inc(x_11); -x_20 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__2(x_19, x_7, x_8, x_9, x_10, x_11, x_12, x_18); -if (lean_obj_tag(x_20) == 0) +lean_dec(x_15); +x_18 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__7___closed__1; +lean_inc(x_10); +x_19 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__2(x_18, x_6, x_7, x_8, x_9, x_10, x_11, x_17); +if (lean_obj_tag(x_19) == 0) { -uint8_t x_21; -x_21 = !lean_is_exclusive(x_20); -if (x_21 == 0) +uint8_t x_20; +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) { -lean_object* x_22; lean_object* x_23; 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; size_t x_31; size_t x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; 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_22 = lean_ctor_get(x_20, 0); -x_23 = lean_ctor_get(x_11, 5); -lean_inc(x_23); -lean_dec(x_11); -x_24 = 0; -x_25 = l_Lean_SourceInfo_fromRef(x_23, x_24); -x_26 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__7; -lean_inc(x_25); -x_27 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_27, 0, x_25); -lean_ctor_set(x_27, 1, x_26); -x_28 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__4___closed__2; -x_29 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_29, 0, x_28); -lean_ctor_set(x_29, 1, x_3); -x_30 = lean_array_get_size(x_4); -x_31 = lean_usize_of_nat(x_30); -lean_dec(x_30); -x_32 = 0; -x_33 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__24(x_5, x_29, x_31, x_32, x_4); +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; size_t x_30; size_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_21 = lean_ctor_get(x_19, 0); +x_22 = lean_ctor_get(x_10, 5); +lean_inc(x_22); +lean_dec(x_10); +x_23 = 0; +x_24 = l_Lean_SourceInfo_fromRef(x_22, x_23); +x_25 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__7; +lean_inc(x_24); +x_26 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +x_27 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__4___closed__2; +x_28 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_2); +x_29 = lean_array_get_size(x_3); +x_30 = lean_usize_of_nat(x_29); lean_dec(x_29); -x_34 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_35 = l_Array_append___rarg(x_34, x_33); -x_36 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; -lean_inc(x_25); -x_37 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_37, 0, x_25); -lean_ctor_set(x_37, 1, x_36); -lean_ctor_set(x_37, 2, x_35); -x_38 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__8; -lean_inc(x_25); -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_25); -lean_ctor_set(x_39, 1, x_38); -lean_inc(x_25); -x_40 = l_Lean_Syntax_node2(x_25, x_36, x_39, x_22); -x_41 = l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__2___closed__5; -lean_inc(x_25); -x_42 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_42, 0, x_25); -lean_ctor_set(x_42, 1, x_41); -x_43 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__8___closed__3; -lean_inc(x_25); -x_44 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_44, 0, x_25); -lean_ctor_set(x_44, 1, x_43); -x_45 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__8___closed__2; -lean_inc(x_25); -x_46 = l_Lean_Syntax_node3(x_25, x_45, x_42, x_44, x_17); -lean_inc(x_25); -x_47 = l_Lean_Syntax_node1(x_25, x_36, x_46); -x_48 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__11; -lean_inc(x_25); -x_49 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_49, 0, x_25); -lean_ctor_set(x_49, 1, x_48); -x_50 = l_Lean_PrettyPrinter_Delaborator_delabForall___lambda__2___closed__3; -x_51 = l_Lean_Syntax_node5(x_25, x_50, x_27, x_37, x_40, x_47, x_49); -lean_ctor_set(x_20, 0, x_51); -return x_20; +x_31 = 0; +x_32 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__8(x_4, x_28, x_30, x_31, x_3); +lean_dec(x_28); +x_33 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; +x_34 = l_Array_append___rarg(x_33, x_32); +x_35 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; +lean_inc(x_24); +x_36 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_36, 0, x_24); +lean_ctor_set(x_36, 1, x_35); +lean_ctor_set(x_36, 2, x_34); +x_37 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__8; +lean_inc(x_24); +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_24); +lean_ctor_set(x_38, 1, x_37); +lean_inc(x_24); +x_39 = l_Lean_Syntax_node2(x_24, x_35, x_38, x_21); +x_40 = l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__2___closed__5; +lean_inc(x_24); +x_41 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_41, 0, x_24); +lean_ctor_set(x_41, 1, x_40); +x_42 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__7___closed__4; +lean_inc(x_24); +x_43 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_43, 0, x_24); +lean_ctor_set(x_43, 1, x_42); +x_44 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__7___closed__3; +lean_inc(x_24); +x_45 = l_Lean_Syntax_node3(x_24, x_44, x_41, x_43, x_16); +lean_inc(x_24); +x_46 = l_Lean_Syntax_node1(x_24, x_35, x_45); +x_47 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__11; +lean_inc(x_24); +x_48 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_48, 0, x_24); +lean_ctor_set(x_48, 1, x_47); +x_49 = l_Lean_PrettyPrinter_Delaborator_delabForall___lambda__2___closed__3; +x_50 = l_Lean_Syntax_node5(x_24, x_49, x_26, x_36, x_39, x_46, x_48); +lean_ctor_set(x_19, 0, x_50); +return x_19; } else { -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; lean_object* x_59; lean_object* x_60; lean_object* x_61; size_t x_62; size_t x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; -x_52 = lean_ctor_get(x_20, 0); -x_53 = lean_ctor_get(x_20, 1); -lean_inc(x_53); +lean_object* x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; size_t x_61; size_t x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_51 = lean_ctor_get(x_19, 0); +x_52 = lean_ctor_get(x_19, 1); lean_inc(x_52); -lean_dec(x_20); -x_54 = lean_ctor_get(x_11, 5); -lean_inc(x_54); -lean_dec(x_11); -x_55 = 0; -x_56 = l_Lean_SourceInfo_fromRef(x_54, x_55); -x_57 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__7; -lean_inc(x_56); -x_58 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_58, 0, x_56); -lean_ctor_set(x_58, 1, x_57); -x_59 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__4___closed__2; -x_60 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_60, 0, x_59); -lean_ctor_set(x_60, 1, x_3); -x_61 = lean_array_get_size(x_4); -x_62 = lean_usize_of_nat(x_61); -lean_dec(x_61); -x_63 = 0; -x_64 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__24(x_5, x_60, x_62, x_63, x_4); +lean_inc(x_51); +lean_dec(x_19); +x_53 = lean_ctor_get(x_10, 5); +lean_inc(x_53); +lean_dec(x_10); +x_54 = 0; +x_55 = l_Lean_SourceInfo_fromRef(x_53, x_54); +x_56 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__7; +lean_inc(x_55); +x_57 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_57, 0, x_55); +lean_ctor_set(x_57, 1, x_56); +x_58 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__4___closed__2; +x_59 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_59, 0, x_58); +lean_ctor_set(x_59, 1, x_2); +x_60 = lean_array_get_size(x_3); +x_61 = lean_usize_of_nat(x_60); lean_dec(x_60); -x_65 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_66 = l_Array_append___rarg(x_65, x_64); -x_67 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; -lean_inc(x_56); -x_68 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_68, 0, x_56); -lean_ctor_set(x_68, 1, x_67); -lean_ctor_set(x_68, 2, x_66); -x_69 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__8; -lean_inc(x_56); -x_70 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_70, 0, x_56); -lean_ctor_set(x_70, 1, x_69); -lean_inc(x_56); -x_71 = l_Lean_Syntax_node2(x_56, x_67, x_70, x_52); -x_72 = l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__2___closed__5; -lean_inc(x_56); -x_73 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_73, 0, x_56); -lean_ctor_set(x_73, 1, x_72); -x_74 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__8___closed__3; -lean_inc(x_56); -x_75 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_75, 0, x_56); -lean_ctor_set(x_75, 1, x_74); -x_76 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__8___closed__2; -lean_inc(x_56); -x_77 = l_Lean_Syntax_node3(x_56, x_76, x_73, x_75, x_17); -lean_inc(x_56); -x_78 = l_Lean_Syntax_node1(x_56, x_67, x_77); -x_79 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__11; -lean_inc(x_56); -x_80 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_80, 0, x_56); -lean_ctor_set(x_80, 1, x_79); -x_81 = l_Lean_PrettyPrinter_Delaborator_delabForall___lambda__2___closed__3; -x_82 = l_Lean_Syntax_node5(x_56, x_81, x_58, x_68, x_71, x_78, x_80); -x_83 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_83, 0, x_82); -lean_ctor_set(x_83, 1, x_53); -return x_83; +x_62 = 0; +x_63 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__8(x_4, x_59, x_61, x_62, x_3); +lean_dec(x_59); +x_64 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; +x_65 = l_Array_append___rarg(x_64, x_63); +x_66 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; +lean_inc(x_55); +x_67 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_67, 0, x_55); +lean_ctor_set(x_67, 1, x_66); +lean_ctor_set(x_67, 2, x_65); +x_68 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__8; +lean_inc(x_55); +x_69 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_69, 0, x_55); +lean_ctor_set(x_69, 1, x_68); +lean_inc(x_55); +x_70 = l_Lean_Syntax_node2(x_55, x_66, x_69, x_51); +x_71 = l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__2___closed__5; +lean_inc(x_55); +x_72 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_72, 0, x_55); +lean_ctor_set(x_72, 1, x_71); +x_73 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__7___closed__4; +lean_inc(x_55); +x_74 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_74, 0, x_55); +lean_ctor_set(x_74, 1, x_73); +x_75 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__7___closed__3; +lean_inc(x_55); +x_76 = l_Lean_Syntax_node3(x_55, x_75, x_72, x_74, x_16); +lean_inc(x_55); +x_77 = l_Lean_Syntax_node1(x_55, x_66, x_76); +x_78 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__11; +lean_inc(x_55); +x_79 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_79, 0, x_55); +lean_ctor_set(x_79, 1, x_78); +x_80 = l_Lean_PrettyPrinter_Delaborator_delabForall___lambda__2___closed__3; +x_81 = l_Lean_Syntax_node5(x_55, x_80, x_57, x_67, x_70, x_77, x_79); +x_82 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_82, 0, x_81); +lean_ctor_set(x_82, 1, x_52); +return x_82; } } else { -uint8_t x_84; -lean_dec(x_17); -lean_dec(x_11); -lean_dec(x_4); +uint8_t x_83; +lean_dec(x_16); +lean_dec(x_10); lean_dec(x_3); -x_84 = !lean_is_exclusive(x_20); -if (x_84 == 0) +lean_dec(x_2); +x_83 = !lean_is_exclusive(x_19); +if (x_83 == 0) { -return x_20; +return x_19; } else { -lean_object* x_85; lean_object* x_86; lean_object* x_87; -x_85 = lean_ctor_get(x_20, 0); -x_86 = lean_ctor_get(x_20, 1); -lean_inc(x_86); +lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_84 = lean_ctor_get(x_19, 0); +x_85 = lean_ctor_get(x_19, 1); lean_inc(x_85); -lean_dec(x_20); -x_87 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_87, 0, x_85); -lean_ctor_set(x_87, 1, x_86); -return x_87; +lean_inc(x_84); +lean_dec(x_19); +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_88; -lean_dec(x_12); +uint8_t x_87; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_4); +lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); -x_88 = !lean_is_exclusive(x_16); -if (x_88 == 0) +x_87 = !lean_is_exclusive(x_15); +if (x_87 == 0) { -return x_16; +return x_15; } else { -lean_object* x_89; lean_object* x_90; lean_object* x_91; -x_89 = lean_ctor_get(x_16, 0); -x_90 = lean_ctor_get(x_16, 1); -lean_inc(x_90); +lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_88 = lean_ctor_get(x_15, 0); +x_89 = lean_ctor_get(x_15, 1); lean_inc(x_89); -lean_dec(x_16); -x_91 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_91, 0, x_89); -lean_ctor_set(x_91, 1, x_90); -return x_91; +lean_inc(x_88); +lean_dec(x_15); +x_90 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_90, 0, x_88); +lean_ctor_set(x_90, 1, x_89); +return x_90; } } } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; 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; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; @@ -41195,7 +41256,7 @@ x_19 = lean_array_get_size(x_2); x_20 = lean_usize_of_nat(x_19); lean_dec(x_19); x_21 = 0; -x_22 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__25(x_3, x_18, x_20, x_21, x_2); +x_22 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__9(x_3, x_18, x_20, x_21, x_2); lean_dec(x_18); x_23 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; x_24 = l_Array_append___rarg(x_23, x_22); @@ -41230,7 +41291,7 @@ lean_ctor_set(x_35, 1, x_11); return x_35; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; 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; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; @@ -41252,7 +41313,7 @@ x_19 = lean_array_get_size(x_2); x_20 = lean_usize_of_nat(x_19); lean_dec(x_19); x_21 = 0; -x_22 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__26(x_3, x_18, x_20, x_21, x_2); +x_22 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__10(x_3, x_18, x_20, x_21, x_2); lean_dec(x_18); x_23 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; x_24 = l_Array_append___rarg(x_23, x_22); @@ -41287,7 +41348,7 @@ lean_ctor_set(x_35, 1, x_11); return x_35; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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, lean_object* x_11) { _start: { lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; 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; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; @@ -41309,7 +41370,7 @@ x_19 = lean_array_get_size(x_2); x_20 = lean_usize_of_nat(x_19); lean_dec(x_19); x_21 = 0; -x_22 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__27(x_3, x_18, x_20, x_21, x_2); +x_22 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__11(x_3, x_18, x_20, x_21, x_2); lean_dec(x_18); x_23 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; x_24 = l_Array_append___rarg(x_23, x_22); @@ -41344,7 +41405,7 @@ lean_ctor_set(x_35, 1, x_11); return x_35; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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, lean_object* x_11) { _start: { lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; 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; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; @@ -41366,7 +41427,7 @@ x_19 = lean_array_get_size(x_2); x_20 = lean_usize_of_nat(x_19); lean_dec(x_19); x_21 = 0; -x_22 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__28(x_3, x_18, x_20, x_21, x_2); +x_22 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__12(x_3, x_18, x_20, x_21, x_2); lean_dec(x_18); x_23 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; x_24 = l_Array_append___rarg(x_23, x_22); @@ -41401,7 +41462,7 @@ lean_ctor_set(x_35, 1, x_11); return x_35; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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, lean_object* x_11) { _start: { lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; 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; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; @@ -41423,7 +41484,7 @@ x_19 = lean_array_get_size(x_2); x_20 = lean_usize_of_nat(x_19); lean_dec(x_19); x_21 = 0; -x_22 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__29(x_3, x_18, x_20, x_21, x_2); +x_22 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__13(x_3, x_18, x_20, x_21, x_2); lean_dec(x_18); x_23 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; x_24 = l_Array_append___rarg(x_23, x_22); @@ -41458,7 +41519,7 @@ lean_ctor_set(x_35, 1, x_11); return x_35; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__14(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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) { _start: { lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; 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; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; @@ -41480,7 +41541,7 @@ x_19 = lean_array_get_size(x_2); x_20 = lean_usize_of_nat(x_19); lean_dec(x_19); x_21 = 0; -x_22 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__30(x_3, x_18, x_20, x_21, x_2); +x_22 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__14(x_3, x_18, x_20, x_21, x_2); lean_dec(x_18); x_23 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; x_24 = l_Array_append___rarg(x_23, x_22); @@ -41515,7 +41576,7 @@ lean_ctor_set(x_35, 1, x_11); return x_35; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__15(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__14(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; 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; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; @@ -41537,7 +41598,7 @@ x_19 = lean_array_get_size(x_2); x_20 = lean_usize_of_nat(x_19); lean_dec(x_19); x_21 = 0; -x_22 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__31(x_3, x_18, x_20, x_21, x_2); +x_22 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__15(x_3, x_18, x_20, x_21, x_2); lean_dec(x_18); x_23 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; x_24 = l_Array_append___rarg(x_23, x_22); @@ -41572,7 +41633,7 @@ lean_ctor_set(x_35, 1, x_11); return x_35; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__16___closed__1() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__15___closed__1() { _start: { lean_object* x_1; @@ -41580,24 +41641,25 @@ x_1 = lean_mk_string_from_bytes("binderDefault", 13); return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__16___closed__2() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__15___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_PrettyPrinter_Delaborator_withTypeAscription___closed__2; x_2 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__3; x_3 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__4; -x_4 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__16___closed__1; +x_4 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__15___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__16(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__15(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_13; -lean_inc(x_10); -x_13 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__1(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_object* x_12; lean_object* x_13; +x_12 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11; +lean_inc(x_9); +x_13 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__1(x_12, x_5, x_6, x_7, x_8, x_9, x_10, x_11); if (lean_obj_tag(x_13) == 0) { uint8_t x_14; @@ -41606,9 +41668,9 @@ if (x_14 == 0) { 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; size_t x_24; size_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; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; 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_15 = lean_ctor_get(x_13, 0); -x_16 = lean_ctor_get(x_10, 5); +x_16 = lean_ctor_get(x_9, 5); lean_inc(x_16); -lean_dec(x_10); +lean_dec(x_9); x_17 = 0; x_18 = l_Lean_SourceInfo_fromRef(x_16, x_17); x_19 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__7; @@ -41619,12 +41681,12 @@ lean_ctor_set(x_20, 1, x_19); x_21 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__4___closed__2; x_22 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_2); -x_23 = lean_array_get_size(x_3); +lean_ctor_set(x_22, 1, x_1); +x_23 = lean_array_get_size(x_2); x_24 = lean_usize_of_nat(x_23); lean_dec(x_23); x_25 = 0; -x_26 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__32(x_4, x_22, x_24, x_25, x_3); +x_26 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__16(x_3, x_22, x_24, x_25, x_2); lean_dec(x_22); x_27 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; x_28 = l_Array_append___rarg(x_27, x_26); @@ -41640,13 +41702,13 @@ x_32 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_32, 0, x_18); lean_ctor_set(x_32, 1, x_31); lean_inc(x_18); -x_33 = l_Lean_Syntax_node2(x_18, x_29, x_32, x_5); +x_33 = l_Lean_Syntax_node2(x_18, x_29, x_32, x_4); x_34 = l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__2___closed__5; lean_inc(x_18); x_35 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_35, 0, x_18); lean_ctor_set(x_35, 1, x_34); -x_36 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__16___closed__2; +x_36 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__15___closed__2; lean_inc(x_18); x_37 = l_Lean_Syntax_node2(x_18, x_36, x_35, x_15); lean_inc(x_18); @@ -41669,9 +41731,9 @@ x_44 = lean_ctor_get(x_13, 1); lean_inc(x_44); lean_inc(x_43); lean_dec(x_13); -x_45 = lean_ctor_get(x_10, 5); +x_45 = lean_ctor_get(x_9, 5); lean_inc(x_45); -lean_dec(x_10); +lean_dec(x_9); x_46 = 0; x_47 = l_Lean_SourceInfo_fromRef(x_45, x_46); x_48 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__7; @@ -41682,12 +41744,12 @@ lean_ctor_set(x_49, 1, x_48); x_50 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__4___closed__2; x_51 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_51, 0, x_50); -lean_ctor_set(x_51, 1, x_2); -x_52 = lean_array_get_size(x_3); +lean_ctor_set(x_51, 1, x_1); +x_52 = lean_array_get_size(x_2); x_53 = lean_usize_of_nat(x_52); lean_dec(x_52); x_54 = 0; -x_55 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__32(x_4, x_51, x_53, x_54, x_3); +x_55 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__16(x_3, x_51, x_53, x_54, x_2); lean_dec(x_51); x_56 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; x_57 = l_Array_append___rarg(x_56, x_55); @@ -41703,13 +41765,13 @@ x_61 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_61, 0, x_47); lean_ctor_set(x_61, 1, x_60); lean_inc(x_47); -x_62 = l_Lean_Syntax_node2(x_47, x_58, x_61, x_5); +x_62 = l_Lean_Syntax_node2(x_47, x_58, x_61, x_4); x_63 = l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__2___closed__5; lean_inc(x_47); x_64 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_64, 0, x_47); lean_ctor_set(x_64, 1, x_63); -x_65 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__16___closed__2; +x_65 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__15___closed__2; lean_inc(x_47); x_66 = l_Lean_Syntax_node2(x_47, x_65, x_64, x_43); lean_inc(x_47); @@ -41730,10 +41792,10 @@ return x_72; else { uint8_t x_73; -lean_dec(x_10); -lean_dec(x_5); -lean_dec(x_3); +lean_dec(x_9); +lean_dec(x_4); lean_dec(x_2); +lean_dec(x_1); x_73 = !lean_is_exclusive(x_13); if (x_73 == 0) { @@ -41755,7 +41817,7 @@ return x_76; } } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__17(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__16(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; 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; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; @@ -41777,7 +41839,7 @@ x_19 = lean_array_get_size(x_2); x_20 = lean_usize_of_nat(x_19); lean_dec(x_19); x_21 = 0; -x_22 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__33(x_3, x_18, x_20, x_21, x_2); +x_22 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__17(x_3, x_18, x_20, x_21, x_2); lean_dec(x_18); x_23 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; x_24 = l_Array_append___rarg(x_23, x_22); @@ -41807,7 +41869,7 @@ lean_ctor_set(x_34, 1, x_11); return x_34; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__18(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__17(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; 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; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; @@ -41829,7 +41891,7 @@ x_19 = lean_array_get_size(x_2); x_20 = lean_usize_of_nat(x_19); lean_dec(x_19); x_21 = 0; -x_22 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__34(x_3, x_18, x_20, x_21, x_2); +x_22 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__18(x_3, x_18, x_20, x_21, x_2); lean_dec(x_18); x_23 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; x_24 = l_Array_append___rarg(x_23, x_22); @@ -41859,10 +41921,10 @@ lean_ctor_set(x_34, 1, x_11); return x_34; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__19(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__18(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -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_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_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_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_10 = lean_ctor_get(x_7, 5); lean_inc(x_10); lean_dec(x_7); @@ -41873,2076 +41935,1013 @@ 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 = lean_box(0); -x_16 = l_Array_back___rarg(x_15, x_1); -x_17 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__8; +x_15 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__8; lean_inc(x_12); -x_18 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_18, 0, x_12); -lean_ctor_set(x_18, 1, x_17); -x_19 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; +x_16 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_16, 0, x_12); +lean_ctor_set(x_16, 1, x_15); +x_17 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; lean_inc(x_12); -x_20 = l_Lean_Syntax_node2(x_12, x_19, x_16, x_18); -x_21 = l_Lean_PrettyPrinter_Delaborator_delabLam___lambda__3___closed__5; +x_18 = l_Lean_Syntax_node2(x_12, x_17, x_1, x_16); +x_19 = l_Lean_PrettyPrinter_Delaborator_delabLam___lambda__3___closed__5; lean_inc(x_12); -x_22 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_22, 0, x_12); -lean_ctor_set(x_22, 1, x_21); -x_23 = l_Lean_PrettyPrinter_Delaborator_delabLam___lambda__3___closed__3; -x_24 = l_Lean_Syntax_node4(x_12, x_23, x_14, x_20, x_2, 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_9); -return x_25; -} -} -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Command", 7); -return x_1; -} -} -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("declSig", 7); -return x_1; -} -} -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_PrettyPrinter_Delaborator_withTypeAscription___closed__2; -x_2 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__3; -x_3 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__1; -x_4 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__4() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("ident", 5); -return x_1; -} -} -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__4; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +x_20 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_20, 0, x_12); +lean_ctor_set(x_20, 1, x_19); +x_21 = l_Lean_PrettyPrinter_Delaborator_delabLam___lambda__3___closed__3; +x_22 = l_Lean_Syntax_node4(x_12, x_21, x_14, x_18, x_2, x_20); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_9); +return x_23; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__6() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("bracketedBinder", 15); +x_1 = lean_mk_string_from_bytes("Lean.PrettyPrinter.Delaborator.delabConstWithSignature.delabParamsAux", 69); return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__7() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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_PrettyPrinter_Delaborator_withTypeAscription___closed__2; -x_2 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__3; -x_3 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__4; -x_4 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__8() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__7; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; -} -} -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__4___closed__2; -x_2 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__10() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5; -x_2 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__1; +x_2 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___closed__1; +x_3 = lean_unsigned_to_nat(1131u); +x_4 = lean_unsigned_to_nat(42u); +x_5 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__3; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___closed__3() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__1), 7, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__6___boxed), 7, 0); return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__12() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__7___boxed), 7, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__13() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11; -x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__1), 8, 1); -lean_closure_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__14() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__13; +x_1 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__7___closed__1; x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__2), 8, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_11; lean_object* x_12; -x_11 = l_Lean_PrettyPrinter_Delaborator_SubExpr_getExpr___at_Lean_PrettyPrinter_Delaborator_getExprKind___spec__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); -switch (lean_obj_tag(x_12)) { -case 0: -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_12); -lean_dec(x_3); -x_13 = lean_ctor_get(x_11, 1); +lean_object* x_12; lean_object* x_13; +x_12 = l_Lean_PrettyPrinter_Delaborator_SubExpr_getExpr___at_Lean_PrettyPrinter_Delaborator_getExprKind___spec__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); -lean_dec(x_11); -lean_inc(x_8); -x_14 = l_Lean_PrettyPrinter_Delaborator_delab(x_4, x_5, x_6, x_7, x_8, x_9, x_13); -if (lean_obj_tag(x_14) == 0) -{ -uint8_t x_15; -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; lean_object* x_20; size_t x_21; size_t x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; size_t x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_16 = lean_ctor_get(x_14, 0); -x_17 = lean_ctor_get(x_8, 5); -lean_inc(x_17); -lean_dec(x_8); -x_18 = 0; -x_19 = l_Lean_SourceInfo_fromRef(x_17, x_18); -x_20 = lean_array_get_size(x_2); -x_21 = lean_usize_of_nat(x_20); -lean_dec(x_20); -x_22 = 0; -x_23 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; -x_24 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__1(x_23, x_21, x_22, x_2); -x_25 = lean_array_get_size(x_24); -x_26 = lean_usize_of_nat(x_25); -lean_dec(x_25); -x_27 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__10; -x_28 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__2(x_27, x_26, x_22, x_24); -x_29 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_30 = l_Array_append___rarg(x_29, x_28); -x_31 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; -lean_inc(x_19); -x_32 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_32, 0, x_19); -lean_ctor_set(x_32, 1, x_31); -lean_ctor_set(x_32, 2, x_30); -x_33 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__8; -lean_inc(x_19); -x_34 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_34, 0, x_19); -lean_ctor_set(x_34, 1, x_33); -x_35 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__9; -lean_inc(x_19); -x_36 = l_Lean_Syntax_node2(x_19, x_35, x_34, x_16); -x_37 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__3; -lean_inc(x_19); -x_38 = l_Lean_Syntax_node2(x_19, x_37, x_32, x_36); -x_39 = l_Lean_PrettyPrinter_Delaborator_declSigWithId_formatter___closed__2; -x_40 = l_Lean_Syntax_node2(x_19, x_39, x_1, x_38); -lean_ctor_set(x_14, 0, x_40); -return x_14; -} -else -{ -lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; lean_object* x_45; lean_object* x_46; size_t x_47; size_t x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; size_t x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_41 = lean_ctor_get(x_14, 0); -x_42 = lean_ctor_get(x_14, 1); -lean_inc(x_42); -lean_inc(x_41); -lean_dec(x_14); -x_43 = lean_ctor_get(x_8, 5); -lean_inc(x_43); -lean_dec(x_8); -x_44 = 0; -x_45 = l_Lean_SourceInfo_fromRef(x_43, x_44); -x_46 = lean_array_get_size(x_2); -x_47 = lean_usize_of_nat(x_46); -lean_dec(x_46); -x_48 = 0; -x_49 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; -x_50 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__1(x_49, x_47, x_48, x_2); -x_51 = lean_array_get_size(x_50); -x_52 = lean_usize_of_nat(x_51); -lean_dec(x_51); -x_53 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__10; -x_54 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__2(x_53, x_52, x_48, x_50); -x_55 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_56 = l_Array_append___rarg(x_55, x_54); -x_57 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; -lean_inc(x_45); -x_58 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_58, 0, x_45); -lean_ctor_set(x_58, 1, x_57); -lean_ctor_set(x_58, 2, x_56); -x_59 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__8; -lean_inc(x_45); -x_60 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_60, 0, x_45); -lean_ctor_set(x_60, 1, x_59); -x_61 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__9; -lean_inc(x_45); -x_62 = l_Lean_Syntax_node2(x_45, x_61, x_60, x_41); -x_63 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__3; -lean_inc(x_45); -x_64 = l_Lean_Syntax_node2(x_45, x_63, x_58, x_62); -x_65 = l_Lean_PrettyPrinter_Delaborator_declSigWithId_formatter___closed__2; -x_66 = l_Lean_Syntax_node2(x_45, x_65, x_1, x_64); -x_67 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_67, 0, x_66); -lean_ctor_set(x_67, 1, x_42); -return x_67; -} -} -else -{ -uint8_t x_68; -lean_dec(x_8); -lean_dec(x_2); -lean_dec(x_1); -x_68 = !lean_is_exclusive(x_14); -if (x_68 == 0) -{ -return x_14; -} -else -{ -lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_69 = lean_ctor_get(x_14, 0); -x_70 = lean_ctor_get(x_14, 1); -lean_inc(x_70); -lean_inc(x_69); -lean_dec(x_14); -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; -} -} -} -case 1: -{ -lean_object* x_72; lean_object* x_73; -lean_dec(x_12); -lean_dec(x_3); -x_72 = lean_ctor_get(x_11, 1); -lean_inc(x_72); -lean_dec(x_11); -lean_inc(x_8); -x_73 = l_Lean_PrettyPrinter_Delaborator_delab(x_4, x_5, x_6, x_7, x_8, x_9, x_72); -if (lean_obj_tag(x_73) == 0) -{ -uint8_t x_74; -x_74 = !lean_is_exclusive(x_73); -if (x_74 == 0) -{ -lean_object* x_75; lean_object* x_76; uint8_t x_77; lean_object* x_78; lean_object* x_79; size_t x_80; size_t x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; size_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; 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_75 = lean_ctor_get(x_73, 0); -x_76 = lean_ctor_get(x_8, 5); -lean_inc(x_76); -lean_dec(x_8); -x_77 = 0; -x_78 = l_Lean_SourceInfo_fromRef(x_76, x_77); -x_79 = lean_array_get_size(x_2); -x_80 = lean_usize_of_nat(x_79); -lean_dec(x_79); -x_81 = 0; -x_82 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; -x_83 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__3(x_82, x_80, x_81, x_2); -x_84 = lean_array_get_size(x_83); -x_85 = lean_usize_of_nat(x_84); -lean_dec(x_84); -x_86 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__10; -x_87 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__4(x_86, x_85, x_81, x_83); -x_88 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_89 = l_Array_append___rarg(x_88, x_87); -x_90 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; -lean_inc(x_78); -x_91 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_91, 0, x_78); -lean_ctor_set(x_91, 1, x_90); -lean_ctor_set(x_91, 2, x_89); -x_92 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__8; -lean_inc(x_78); -x_93 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_93, 0, x_78); -lean_ctor_set(x_93, 1, x_92); -x_94 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__9; -lean_inc(x_78); -x_95 = l_Lean_Syntax_node2(x_78, x_94, x_93, x_75); -x_96 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__3; -lean_inc(x_78); -x_97 = l_Lean_Syntax_node2(x_78, x_96, x_91, x_95); -x_98 = l_Lean_PrettyPrinter_Delaborator_declSigWithId_formatter___closed__2; -x_99 = l_Lean_Syntax_node2(x_78, x_98, x_1, x_97); -lean_ctor_set(x_73, 0, x_99); -return x_73; -} -else -{ -lean_object* x_100; lean_object* x_101; lean_object* x_102; uint8_t x_103; lean_object* x_104; lean_object* x_105; size_t x_106; size_t x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; size_t 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; -x_100 = lean_ctor_get(x_73, 0); -x_101 = lean_ctor_get(x_73, 1); -lean_inc(x_101); -lean_inc(x_100); -lean_dec(x_73); -x_102 = lean_ctor_get(x_8, 5); -lean_inc(x_102); -lean_dec(x_8); -x_103 = 0; -x_104 = l_Lean_SourceInfo_fromRef(x_102, x_103); -x_105 = lean_array_get_size(x_2); -x_106 = lean_usize_of_nat(x_105); -lean_dec(x_105); -x_107 = 0; -x_108 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; -x_109 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__3(x_108, x_106, x_107, x_2); -x_110 = lean_array_get_size(x_109); -x_111 = lean_usize_of_nat(x_110); -lean_dec(x_110); -x_112 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__10; -x_113 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__4(x_112, x_111, x_107, x_109); -x_114 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_115 = l_Array_append___rarg(x_114, x_113); -x_116 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; -lean_inc(x_104); -x_117 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_117, 0, x_104); -lean_ctor_set(x_117, 1, x_116); -lean_ctor_set(x_117, 2, x_115); -x_118 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__8; -lean_inc(x_104); -x_119 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_119, 0, x_104); -lean_ctor_set(x_119, 1, x_118); -x_120 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__9; -lean_inc(x_104); -x_121 = l_Lean_Syntax_node2(x_104, x_120, x_119, x_100); -x_122 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__3; -lean_inc(x_104); -x_123 = l_Lean_Syntax_node2(x_104, x_122, x_117, x_121); -x_124 = l_Lean_PrettyPrinter_Delaborator_declSigWithId_formatter___closed__2; -x_125 = l_Lean_Syntax_node2(x_104, x_124, x_1, x_123); -x_126 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_126, 0, x_125); -lean_ctor_set(x_126, 1, x_101); -return x_126; -} -} -else -{ -uint8_t x_127; -lean_dec(x_8); -lean_dec(x_2); -lean_dec(x_1); -x_127 = !lean_is_exclusive(x_73); -if (x_127 == 0) -{ -return x_73; -} -else -{ -lean_object* x_128; lean_object* x_129; lean_object* x_130; -x_128 = lean_ctor_get(x_73, 0); -x_129 = lean_ctor_get(x_73, 1); -lean_inc(x_129); -lean_inc(x_128); -lean_dec(x_73); -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; -} -} -} -case 2: -{ -lean_object* x_131; lean_object* x_132; -lean_dec(x_12); -lean_dec(x_3); -x_131 = lean_ctor_get(x_11, 1); -lean_inc(x_131); -lean_dec(x_11); -lean_inc(x_8); -x_132 = l_Lean_PrettyPrinter_Delaborator_delab(x_4, x_5, x_6, x_7, x_8, x_9, x_131); -if (lean_obj_tag(x_132) == 0) -{ -uint8_t x_133; -x_133 = !lean_is_exclusive(x_132); -if (x_133 == 0) -{ -lean_object* x_134; lean_object* x_135; uint8_t x_136; lean_object* x_137; lean_object* x_138; size_t x_139; size_t x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; size_t x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; -x_134 = lean_ctor_get(x_132, 0); -x_135 = lean_ctor_get(x_8, 5); -lean_inc(x_135); -lean_dec(x_8); -x_136 = 0; -x_137 = l_Lean_SourceInfo_fromRef(x_135, x_136); -x_138 = lean_array_get_size(x_2); -x_139 = lean_usize_of_nat(x_138); -lean_dec(x_138); -x_140 = 0; -x_141 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; -x_142 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__5(x_141, x_139, x_140, x_2); -x_143 = lean_array_get_size(x_142); -x_144 = lean_usize_of_nat(x_143); -lean_dec(x_143); -x_145 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__10; -x_146 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__6(x_145, x_144, x_140, x_142); -x_147 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_148 = l_Array_append___rarg(x_147, x_146); -x_149 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; -lean_inc(x_137); -x_150 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_150, 0, x_137); -lean_ctor_set(x_150, 1, x_149); -lean_ctor_set(x_150, 2, x_148); -x_151 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__8; -lean_inc(x_137); -x_152 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_152, 0, x_137); -lean_ctor_set(x_152, 1, x_151); -x_153 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__9; -lean_inc(x_137); -x_154 = l_Lean_Syntax_node2(x_137, x_153, x_152, x_134); -x_155 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__3; -lean_inc(x_137); -x_156 = l_Lean_Syntax_node2(x_137, x_155, x_150, x_154); -x_157 = l_Lean_PrettyPrinter_Delaborator_declSigWithId_formatter___closed__2; -x_158 = l_Lean_Syntax_node2(x_137, x_157, x_1, x_156); -lean_ctor_set(x_132, 0, x_158); -return x_132; -} -else -{ -lean_object* x_159; lean_object* x_160; lean_object* x_161; uint8_t x_162; lean_object* x_163; lean_object* x_164; size_t x_165; size_t x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; size_t 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; -x_159 = lean_ctor_get(x_132, 0); -x_160 = lean_ctor_get(x_132, 1); -lean_inc(x_160); -lean_inc(x_159); -lean_dec(x_132); -x_161 = lean_ctor_get(x_8, 5); -lean_inc(x_161); -lean_dec(x_8); -x_162 = 0; -x_163 = l_Lean_SourceInfo_fromRef(x_161, x_162); -x_164 = lean_array_get_size(x_2); -x_165 = lean_usize_of_nat(x_164); -lean_dec(x_164); -x_166 = 0; -x_167 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; -x_168 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__5(x_167, x_165, x_166, x_2); -x_169 = lean_array_get_size(x_168); -x_170 = lean_usize_of_nat(x_169); -lean_dec(x_169); -x_171 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__10; -x_172 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__6(x_171, x_170, x_166, x_168); -x_173 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_174 = l_Array_append___rarg(x_173, x_172); -x_175 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; -lean_inc(x_163); -x_176 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_176, 0, x_163); -lean_ctor_set(x_176, 1, x_175); -lean_ctor_set(x_176, 2, x_174); -x_177 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__8; -lean_inc(x_163); -x_178 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_178, 0, x_163); -lean_ctor_set(x_178, 1, x_177); -x_179 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__9; -lean_inc(x_163); -x_180 = l_Lean_Syntax_node2(x_163, x_179, x_178, x_159); -x_181 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__3; -lean_inc(x_163); -x_182 = l_Lean_Syntax_node2(x_163, x_181, x_176, x_180); -x_183 = l_Lean_PrettyPrinter_Delaborator_declSigWithId_formatter___closed__2; -x_184 = l_Lean_Syntax_node2(x_163, x_183, x_1, x_182); -x_185 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_185, 0, x_184); -lean_ctor_set(x_185, 1, x_160); -return x_185; -} -} -else -{ -uint8_t x_186; -lean_dec(x_8); -lean_dec(x_2); -lean_dec(x_1); -x_186 = !lean_is_exclusive(x_132); -if (x_186 == 0) -{ -return x_132; -} -else -{ -lean_object* x_187; lean_object* x_188; lean_object* x_189; -x_187 = lean_ctor_get(x_132, 0); -x_188 = lean_ctor_get(x_132, 1); -lean_inc(x_188); -lean_inc(x_187); -lean_dec(x_132); -x_189 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_189, 0, x_187); -lean_ctor_set(x_189, 1, x_188); -return x_189; -} -} -} -case 3: -{ -lean_object* x_190; lean_object* x_191; -lean_dec(x_12); -lean_dec(x_3); -x_190 = lean_ctor_get(x_11, 1); -lean_inc(x_190); -lean_dec(x_11); -lean_inc(x_8); -x_191 = l_Lean_PrettyPrinter_Delaborator_delab(x_4, x_5, x_6, x_7, x_8, x_9, x_190); -if (lean_obj_tag(x_191) == 0) -{ -uint8_t x_192; -x_192 = !lean_is_exclusive(x_191); -if (x_192 == 0) -{ -lean_object* x_193; lean_object* x_194; uint8_t x_195; lean_object* x_196; lean_object* x_197; size_t x_198; size_t x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; size_t x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; -x_193 = lean_ctor_get(x_191, 0); -x_194 = lean_ctor_get(x_8, 5); -lean_inc(x_194); -lean_dec(x_8); -x_195 = 0; -x_196 = l_Lean_SourceInfo_fromRef(x_194, x_195); -x_197 = lean_array_get_size(x_2); -x_198 = lean_usize_of_nat(x_197); -lean_dec(x_197); -x_199 = 0; -x_200 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; -x_201 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__7(x_200, x_198, x_199, x_2); -x_202 = lean_array_get_size(x_201); -x_203 = lean_usize_of_nat(x_202); -lean_dec(x_202); -x_204 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__10; -x_205 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__8(x_204, x_203, x_199, x_201); -x_206 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_207 = l_Array_append___rarg(x_206, x_205); -x_208 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; -lean_inc(x_196); -x_209 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_209, 0, x_196); -lean_ctor_set(x_209, 1, x_208); -lean_ctor_set(x_209, 2, x_207); -x_210 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__8; -lean_inc(x_196); -x_211 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_211, 0, x_196); -lean_ctor_set(x_211, 1, x_210); -x_212 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__9; -lean_inc(x_196); -x_213 = l_Lean_Syntax_node2(x_196, x_212, x_211, x_193); -x_214 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__3; -lean_inc(x_196); -x_215 = l_Lean_Syntax_node2(x_196, x_214, x_209, x_213); -x_216 = l_Lean_PrettyPrinter_Delaborator_declSigWithId_formatter___closed__2; -x_217 = l_Lean_Syntax_node2(x_196, x_216, x_1, x_215); -lean_ctor_set(x_191, 0, x_217); -return x_191; -} -else -{ -lean_object* x_218; lean_object* x_219; lean_object* x_220; uint8_t x_221; lean_object* x_222; lean_object* x_223; size_t x_224; size_t x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; size_t 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; -x_218 = lean_ctor_get(x_191, 0); -x_219 = lean_ctor_get(x_191, 1); -lean_inc(x_219); -lean_inc(x_218); -lean_dec(x_191); -x_220 = lean_ctor_get(x_8, 5); -lean_inc(x_220); -lean_dec(x_8); -x_221 = 0; -x_222 = l_Lean_SourceInfo_fromRef(x_220, x_221); -x_223 = lean_array_get_size(x_2); -x_224 = lean_usize_of_nat(x_223); -lean_dec(x_223); -x_225 = 0; -x_226 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; -x_227 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__7(x_226, x_224, x_225, x_2); -x_228 = lean_array_get_size(x_227); -x_229 = lean_usize_of_nat(x_228); -lean_dec(x_228); -x_230 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__10; -x_231 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__8(x_230, x_229, x_225, x_227); -x_232 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_233 = l_Array_append___rarg(x_232, x_231); -x_234 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; -lean_inc(x_222); -x_235 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_235, 0, x_222); -lean_ctor_set(x_235, 1, x_234); -lean_ctor_set(x_235, 2, x_233); -x_236 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__8; -lean_inc(x_222); -x_237 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_237, 0, x_222); -lean_ctor_set(x_237, 1, x_236); -x_238 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__9; -lean_inc(x_222); -x_239 = l_Lean_Syntax_node2(x_222, x_238, x_237, x_218); -x_240 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__3; -lean_inc(x_222); -x_241 = l_Lean_Syntax_node2(x_222, x_240, x_235, x_239); -x_242 = l_Lean_PrettyPrinter_Delaborator_declSigWithId_formatter___closed__2; -x_243 = l_Lean_Syntax_node2(x_222, x_242, x_1, x_241); -x_244 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_244, 0, x_243); -lean_ctor_set(x_244, 1, x_219); -return x_244; -} -} -else -{ -uint8_t x_245; -lean_dec(x_8); -lean_dec(x_2); -lean_dec(x_1); -x_245 = !lean_is_exclusive(x_191); -if (x_245 == 0) -{ -return x_191; -} -else +if (lean_obj_tag(x_13) == 7) { -lean_object* x_246; lean_object* x_247; lean_object* x_248; -x_246 = lean_ctor_get(x_191, 0); -x_247 = lean_ctor_get(x_191, 1); -lean_inc(x_247); -lean_inc(x_246); -lean_dec(x_191); -x_248 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_248, 0, x_246); -lean_ctor_set(x_248, 1, x_247); -return x_248; -} -} -} -case 4: -{ -lean_object* x_249; lean_object* x_250; -lean_dec(x_12); -lean_dec(x_3); -x_249 = lean_ctor_get(x_11, 1); -lean_inc(x_249); -lean_dec(x_11); -lean_inc(x_8); -x_250 = l_Lean_PrettyPrinter_Delaborator_delab(x_4, x_5, x_6, x_7, x_8, x_9, x_249); -if (lean_obj_tag(x_250) == 0) -{ -uint8_t x_251; -x_251 = !lean_is_exclusive(x_250); -if (x_251 == 0) -{ -lean_object* x_252; lean_object* x_253; uint8_t x_254; lean_object* x_255; lean_object* x_256; size_t x_257; size_t x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; size_t x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_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_252 = lean_ctor_get(x_250, 0); -x_253 = lean_ctor_get(x_8, 5); -lean_inc(x_253); -lean_dec(x_8); -x_254 = 0; -x_255 = l_Lean_SourceInfo_fromRef(x_253, x_254); -x_256 = lean_array_get_size(x_2); -x_257 = lean_usize_of_nat(x_256); -lean_dec(x_256); -x_258 = 0; -x_259 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; -x_260 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__9(x_259, x_257, x_258, x_2); -x_261 = lean_array_get_size(x_260); -x_262 = lean_usize_of_nat(x_261); -lean_dec(x_261); -x_263 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__10; -x_264 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__10(x_263, x_262, x_258, x_260); -x_265 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_266 = l_Array_append___rarg(x_265, x_264); -x_267 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; -lean_inc(x_255); -x_268 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_268, 0, x_255); -lean_ctor_set(x_268, 1, x_267); -lean_ctor_set(x_268, 2, x_266); -x_269 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__8; -lean_inc(x_255); -x_270 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_270, 0, x_255); -lean_ctor_set(x_270, 1, x_269); -x_271 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__9; -lean_inc(x_255); -x_272 = l_Lean_Syntax_node2(x_255, x_271, x_270, x_252); -x_273 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__3; -lean_inc(x_255); -x_274 = l_Lean_Syntax_node2(x_255, x_273, x_268, x_272); -x_275 = l_Lean_PrettyPrinter_Delaborator_declSigWithId_formatter___closed__2; -x_276 = l_Lean_Syntax_node2(x_255, x_275, x_1, x_274); -lean_ctor_set(x_250, 0, x_276); -return x_250; -} -else -{ -lean_object* x_277; lean_object* x_278; lean_object* x_279; uint8_t x_280; lean_object* x_281; lean_object* x_282; size_t x_283; size_t x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; size_t 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_277 = lean_ctor_get(x_250, 0); -x_278 = lean_ctor_get(x_250, 1); -lean_inc(x_278); -lean_inc(x_277); -lean_dec(x_250); -x_279 = lean_ctor_get(x_8, 5); -lean_inc(x_279); -lean_dec(x_8); -x_280 = 0; -x_281 = l_Lean_SourceInfo_fromRef(x_279, x_280); -x_282 = lean_array_get_size(x_2); -x_283 = lean_usize_of_nat(x_282); -lean_dec(x_282); -x_284 = 0; -x_285 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; -x_286 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__9(x_285, x_283, x_284, x_2); -x_287 = lean_array_get_size(x_286); -x_288 = lean_usize_of_nat(x_287); -lean_dec(x_287); -x_289 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__10; -x_290 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__10(x_289, x_288, x_284, x_286); -x_291 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_292 = l_Array_append___rarg(x_291, x_290); -x_293 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; -lean_inc(x_281); -x_294 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_294, 0, x_281); -lean_ctor_set(x_294, 1, x_293); -lean_ctor_set(x_294, 2, x_292); -x_295 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__8; -lean_inc(x_281); -x_296 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_296, 0, x_281); -lean_ctor_set(x_296, 1, x_295); -x_297 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__9; -lean_inc(x_281); -x_298 = l_Lean_Syntax_node2(x_281, x_297, x_296, x_277); -x_299 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__3; -lean_inc(x_281); -x_300 = l_Lean_Syntax_node2(x_281, x_299, x_294, x_298); -x_301 = l_Lean_PrettyPrinter_Delaborator_declSigWithId_formatter___closed__2; -x_302 = l_Lean_Syntax_node2(x_281, x_301, x_1, x_300); -x_303 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_303, 0, x_302); -lean_ctor_set(x_303, 1, x_278); -return x_303; -} -} -else -{ -uint8_t x_304; -lean_dec(x_8); -lean_dec(x_2); -lean_dec(x_1); -x_304 = !lean_is_exclusive(x_250); -if (x_304 == 0) -{ -return x_250; -} -else -{ -lean_object* x_305; lean_object* x_306; lean_object* x_307; -x_305 = lean_ctor_get(x_250, 0); -x_306 = lean_ctor_get(x_250, 1); -lean_inc(x_306); -lean_inc(x_305); -lean_dec(x_250); -x_307 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_307, 0, x_305); -lean_ctor_set(x_307, 1, x_306); -return x_307; -} -} -} -case 5: -{ -lean_object* x_308; lean_object* x_309; -lean_dec(x_12); -lean_dec(x_3); -x_308 = lean_ctor_get(x_11, 1); -lean_inc(x_308); -lean_dec(x_11); -lean_inc(x_8); -x_309 = l_Lean_PrettyPrinter_Delaborator_delab(x_4, x_5, x_6, x_7, x_8, x_9, x_308); -if (lean_obj_tag(x_309) == 0) -{ -uint8_t x_310; -x_310 = !lean_is_exclusive(x_309); -if (x_310 == 0) -{ -lean_object* x_311; lean_object* x_312; uint8_t x_313; lean_object* x_314; lean_object* x_315; size_t x_316; size_t x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; size_t 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; -x_311 = lean_ctor_get(x_309, 0); -x_312 = lean_ctor_get(x_8, 5); -lean_inc(x_312); -lean_dec(x_8); -x_313 = 0; -x_314 = l_Lean_SourceInfo_fromRef(x_312, x_313); -x_315 = lean_array_get_size(x_2); -x_316 = lean_usize_of_nat(x_315); -lean_dec(x_315); -x_317 = 0; -x_318 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; -x_319 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__11(x_318, x_316, x_317, x_2); -x_320 = lean_array_get_size(x_319); -x_321 = lean_usize_of_nat(x_320); -lean_dec(x_320); -x_322 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__10; -x_323 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__12(x_322, x_321, x_317, x_319); -x_324 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_325 = l_Array_append___rarg(x_324, x_323); -x_326 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; -lean_inc(x_314); -x_327 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_327, 0, x_314); -lean_ctor_set(x_327, 1, x_326); -lean_ctor_set(x_327, 2, x_325); -x_328 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__8; -lean_inc(x_314); -x_329 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_329, 0, x_314); -lean_ctor_set(x_329, 1, x_328); -x_330 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__9; -lean_inc(x_314); -x_331 = l_Lean_Syntax_node2(x_314, x_330, x_329, x_311); -x_332 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__3; -lean_inc(x_314); -x_333 = l_Lean_Syntax_node2(x_314, x_332, x_327, x_331); -x_334 = l_Lean_PrettyPrinter_Delaborator_declSigWithId_formatter___closed__2; -x_335 = l_Lean_Syntax_node2(x_314, x_334, x_1, x_333); -lean_ctor_set(x_309, 0, x_335); -return x_309; -} -else -{ -lean_object* x_336; lean_object* x_337; lean_object* x_338; uint8_t x_339; lean_object* x_340; lean_object* x_341; size_t x_342; size_t x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; size_t 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; -x_336 = lean_ctor_get(x_309, 0); -x_337 = lean_ctor_get(x_309, 1); -lean_inc(x_337); -lean_inc(x_336); -lean_dec(x_309); -x_338 = lean_ctor_get(x_8, 5); -lean_inc(x_338); -lean_dec(x_8); -x_339 = 0; -x_340 = l_Lean_SourceInfo_fromRef(x_338, x_339); -x_341 = lean_array_get_size(x_2); -x_342 = lean_usize_of_nat(x_341); -lean_dec(x_341); -x_343 = 0; -x_344 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; -x_345 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__11(x_344, x_342, x_343, x_2); -x_346 = lean_array_get_size(x_345); -x_347 = lean_usize_of_nat(x_346); -lean_dec(x_346); -x_348 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__10; -x_349 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__12(x_348, x_347, x_343, x_345); -x_350 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_351 = l_Array_append___rarg(x_350, x_349); -x_352 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; -lean_inc(x_340); -x_353 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_353, 0, x_340); -lean_ctor_set(x_353, 1, x_352); -lean_ctor_set(x_353, 2, x_351); -x_354 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__8; -lean_inc(x_340); -x_355 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_355, 0, x_340); -lean_ctor_set(x_355, 1, x_354); -x_356 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__9; -lean_inc(x_340); -x_357 = l_Lean_Syntax_node2(x_340, x_356, x_355, x_336); -x_358 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__3; -lean_inc(x_340); -x_359 = l_Lean_Syntax_node2(x_340, x_358, x_353, x_357); -x_360 = l_Lean_PrettyPrinter_Delaborator_declSigWithId_formatter___closed__2; -x_361 = l_Lean_Syntax_node2(x_340, x_360, x_1, x_359); -x_362 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_362, 0, x_361); -lean_ctor_set(x_362, 1, x_337); -return x_362; -} -} -else -{ -uint8_t x_363; -lean_dec(x_8); -lean_dec(x_2); -lean_dec(x_1); -x_363 = !lean_is_exclusive(x_309); -if (x_363 == 0) -{ -return x_309; -} -else -{ -lean_object* x_364; lean_object* x_365; lean_object* x_366; -x_364 = lean_ctor_get(x_309, 0); -x_365 = lean_ctor_get(x_309, 1); -lean_inc(x_365); -lean_inc(x_364); -lean_dec(x_309); -x_366 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_366, 0, x_364); -lean_ctor_set(x_366, 1, x_365); -return x_366; -} -} -} -case 6: -{ -lean_object* x_367; lean_object* x_368; +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; lean_object* x_23; uint8_t x_24; +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); lean_dec(x_12); -lean_dec(x_3); -x_367 = lean_ctor_get(x_11, 1); -lean_inc(x_367); -lean_dec(x_11); -lean_inc(x_8); -x_368 = l_Lean_PrettyPrinter_Delaborator_delab(x_4, x_5, x_6, x_7, x_8, x_9, x_367); -if (lean_obj_tag(x_368) == 0) -{ -uint8_t x_369; -x_369 = !lean_is_exclusive(x_368); -if (x_369 == 0) -{ -lean_object* x_370; lean_object* x_371; uint8_t x_372; lean_object* x_373; lean_object* x_374; size_t x_375; size_t x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; size_t x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; -x_370 = lean_ctor_get(x_368, 0); -x_371 = lean_ctor_get(x_8, 5); -lean_inc(x_371); -lean_dec(x_8); -x_372 = 0; -x_373 = l_Lean_SourceInfo_fromRef(x_371, x_372); -x_374 = lean_array_get_size(x_2); -x_375 = lean_usize_of_nat(x_374); -lean_dec(x_374); -x_376 = 0; -x_377 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; -x_378 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__13(x_377, x_375, x_376, x_2); -x_379 = lean_array_get_size(x_378); -x_380 = lean_usize_of_nat(x_379); -lean_dec(x_379); -x_381 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__10; -x_382 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__14(x_381, x_380, x_376, x_378); -x_383 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_384 = l_Array_append___rarg(x_383, x_382); -x_385 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; -lean_inc(x_373); -x_386 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_386, 0, x_373); -lean_ctor_set(x_386, 1, x_385); -lean_ctor_set(x_386, 2, x_384); -x_387 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__8; -lean_inc(x_373); -x_388 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_388, 0, x_373); -lean_ctor_set(x_388, 1, x_387); -x_389 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__9; -lean_inc(x_373); -x_390 = l_Lean_Syntax_node2(x_373, x_389, x_388, x_370); -x_391 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__3; -lean_inc(x_373); -x_392 = l_Lean_Syntax_node2(x_373, x_391, x_386, x_390); -x_393 = l_Lean_PrettyPrinter_Delaborator_declSigWithId_formatter___closed__2; -x_394 = l_Lean_Syntax_node2(x_373, x_393, x_1, x_392); -lean_ctor_set(x_368, 0, x_394); -return x_368; -} -else -{ -lean_object* x_395; lean_object* x_396; lean_object* x_397; uint8_t x_398; lean_object* x_399; lean_object* x_400; size_t x_401; size_t x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; size_t 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; -x_395 = lean_ctor_get(x_368, 0); -x_396 = lean_ctor_get(x_368, 1); -lean_inc(x_396); -lean_inc(x_395); -lean_dec(x_368); -x_397 = lean_ctor_get(x_8, 5); -lean_inc(x_397); -lean_dec(x_8); -x_398 = 0; -x_399 = l_Lean_SourceInfo_fromRef(x_397, x_398); -x_400 = lean_array_get_size(x_2); -x_401 = lean_usize_of_nat(x_400); -lean_dec(x_400); -x_402 = 0; -x_403 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; -x_404 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__13(x_403, x_401, x_402, x_2); -x_405 = lean_array_get_size(x_404); -x_406 = lean_usize_of_nat(x_405); -lean_dec(x_405); -x_407 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__10; -x_408 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__14(x_407, x_406, x_402, x_404); -x_409 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_410 = l_Array_append___rarg(x_409, x_408); -x_411 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; -lean_inc(x_399); -x_412 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_412, 0, x_399); -lean_ctor_set(x_412, 1, x_411); -lean_ctor_set(x_412, 2, x_410); -x_413 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__8; -lean_inc(x_399); -x_414 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_414, 0, x_399); -lean_ctor_set(x_414, 1, x_413); -x_415 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__9; -lean_inc(x_399); -x_416 = l_Lean_Syntax_node2(x_399, x_415, x_414, x_395); -x_417 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__3; -lean_inc(x_399); -x_418 = l_Lean_Syntax_node2(x_399, x_417, x_412, x_416); -x_419 = l_Lean_PrettyPrinter_Delaborator_declSigWithId_formatter___closed__2; -x_420 = l_Lean_Syntax_node2(x_399, x_419, x_1, x_418); -x_421 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_421, 0, x_420); -lean_ctor_set(x_421, 1, x_396); -return x_421; -} -} -else -{ -uint8_t x_422; -lean_dec(x_8); -lean_dec(x_2); -lean_dec(x_1); -x_422 = !lean_is_exclusive(x_368); -if (x_422 == 0) -{ -return x_368; -} -else -{ -lean_object* x_423; lean_object* x_424; lean_object* x_425; -x_423 = lean_ctor_get(x_368, 0); -x_424 = lean_ctor_get(x_368, 1); -lean_inc(x_424); -lean_inc(x_423); -lean_dec(x_368); -x_425 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_425, 0, x_423); -lean_ctor_set(x_425, 1, x_424); -return x_425; -} -} -} -case 7: +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); +x_18 = lean_ctor_get_uint8(x_13, sizeof(void*)*3 + 8); +x_19 = lean_box(0); +lean_inc(x_15); +x_20 = l_Lean_RBNode_insert___at_Lean_NameSet_insert___spec__1(x_1, x_15, x_19); +lean_inc(x_15); +x_21 = lean_mk_syntax_ident(x_15); +x_22 = lean_box(0); +lean_inc(x_21); +x_23 = lean_array_push(x_4, x_21); +x_24 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_shouldGroupWithNext(x_20, x_13, x_17); +lean_dec(x_17); +lean_dec(x_13); +if (x_24 == 0) { -lean_object* x_426; lean_object* x_427; lean_object* x_428; uint8_t x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; -x_426 = lean_ctor_get(x_11, 1); -lean_inc(x_426); -lean_dec(x_11); -x_427 = lean_ctor_get(x_12, 0); -lean_inc(x_427); -x_428 = lean_ctor_get(x_12, 1); -lean_inc(x_428); -x_429 = lean_ctor_get_uint8(x_12, sizeof(void*)*3 + 8); -lean_dec(x_12); -x_430 = lean_box(0); -lean_inc(x_427); -x_431 = lean_mk_syntax_ident(x_427); -x_432 = l_Lean_PrettyPrinter_Delaborator_annotateCurPos(x_431, x_4, x_5, x_6, x_7, x_8, x_9, x_426); -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_array_push(x_3, x_433); -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_436 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_shouldGroupWithNext(x_4, x_5, x_6, x_7, x_8, x_9, x_434); -if (lean_obj_tag(x_436) == 0) +switch (x_18) { +case 0: { -lean_object* x_437; uint8_t x_438; -x_437 = lean_ctor_get(x_436, 0); -lean_inc(x_437); -x_438 = lean_unbox(x_437); -lean_dec(x_437); -if (x_438 == 0) +uint8_t x_25; +lean_dec(x_21); +x_25 = l_Lean_Expr_isOptParam(x_16); +if (x_25 == 0) { -switch (x_429) { -case 0: +lean_object* x_26; +x_26 = l_Lean_Expr_getAutoParamTactic_x3f(x_16); +lean_dec(x_16); +if (lean_obj_tag(x_26) == 0) { -lean_object* x_439; uint8_t x_440; -x_439 = lean_ctor_get(x_436, 1); -lean_inc(x_439); -lean_dec(x_436); -x_440 = l_Lean_Expr_isOptParam(x_428); -if (x_440 == 0) -{ -lean_object* x_441; -x_441 = l_Lean_Expr_getAutoParamTactic_x3f(x_428); -lean_dec(x_428); -if (lean_obj_tag(x_441) == 0) -{ -lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; -x_442 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5; -x_443 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__2___boxed), 11, 3); -lean_closure_set(x_443, 0, x_430); -lean_closure_set(x_443, 1, x_435); -lean_closure_set(x_443, 2, x_442); -x_444 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11; -x_445 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__3___rarg), 9, 2); -lean_closure_set(x_445, 0, x_444); -lean_closure_set(x_445, 1, x_443); +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_27 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5; +x_28 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__1___boxed), 11, 3); +lean_closure_set(x_28, 0, x_22); +lean_closure_set(x_28, 1, x_23); +lean_closure_set(x_28, 2, x_27); +x_29 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11; +x_30 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__3___rarg), 9, 2); +lean_closure_set(x_30, 0, x_29); +lean_closure_set(x_30, 1, x_28); +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_446 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingDomain___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__16(x_445, x_4, x_5, x_6, x_7, x_8, x_9, x_439); -if (lean_obj_tag(x_446) == 0) +x_31 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingDomain___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__3(x_30, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +if (lean_obj_tag(x_31) == 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; -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_449 = lean_array_push(x_2, x_447); -x_450 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_451 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams), 10, 3); -lean_closure_set(x_451, 0, x_1); -lean_closure_set(x_451, 1, x_449); -lean_closure_set(x_451, 2, x_450); -x_452 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(x_427, x_451, x_4, x_5, x_6, x_7, x_8, x_9, x_448); -return x_452; +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +x_33 = lean_ctor_get(x_31, 1); +lean_inc(x_33); +lean_dec(x_31); +x_34 = lean_array_push(x_3, x_32); +x_35 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams), 10, 3); +lean_closure_set(x_35, 0, x_20); +lean_closure_set(x_35, 1, x_2); +lean_closure_set(x_35, 2, x_34); +x_36 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(x_15, x_35, x_5, x_6, x_7, x_8, x_9, x_10, x_33); +return x_36; } else { -uint8_t x_453; -lean_dec(x_427); +uint8_t x_37; +lean_dec(x_20); +lean_dec(x_15); +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_453 = !lean_is_exclusive(x_446); -if (x_453 == 0) +x_37 = !lean_is_exclusive(x_31); +if (x_37 == 0) { -return x_446; +return x_31; } else { -lean_object* x_454; lean_object* x_455; lean_object* x_456; -x_454 = lean_ctor_get(x_446, 0); -x_455 = lean_ctor_get(x_446, 1); -lean_inc(x_455); -lean_inc(x_454); -lean_dec(x_446); -x_456 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_456, 0, x_454); -lean_ctor_set(x_456, 1, x_455); -return x_456; +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_31, 0); +x_39 = lean_ctor_get(x_31, 1); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_31); +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_457; -x_457 = lean_ctor_get(x_441, 0); -lean_inc(x_457); -lean_dec(x_441); -switch (lean_obj_tag(x_457)) { +lean_object* x_41; +x_41 = lean_ctor_get(x_26, 0); +lean_inc(x_41); +lean_dec(x_26); +switch (lean_obj_tag(x_41)) { case 0: { -lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; -lean_dec(x_457); -x_458 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5; -x_459 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__3___boxed), 11, 3); -lean_closure_set(x_459, 0, x_430); -lean_closure_set(x_459, 1, x_435); -lean_closure_set(x_459, 2, x_458); -x_460 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11; -x_461 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__3___rarg), 9, 2); -lean_closure_set(x_461, 0, x_460); -lean_closure_set(x_461, 1, x_459); +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +lean_dec(x_41); +x_42 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5; +x_43 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__2___boxed), 11, 3); +lean_closure_set(x_43, 0, x_22); +lean_closure_set(x_43, 1, x_23); +lean_closure_set(x_43, 2, x_42); +x_44 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11; +x_45 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__3___rarg), 9, 2); +lean_closure_set(x_45, 0, x_44); +lean_closure_set(x_45, 1, x_43); +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_462 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingDomain___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__16(x_461, x_4, x_5, x_6, x_7, x_8, x_9, x_439); -if (lean_obj_tag(x_462) == 0) +x_46 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingDomain___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__3(x_45, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +if (lean_obj_tag(x_46) == 0) { -lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; -x_463 = lean_ctor_get(x_462, 0); -lean_inc(x_463); -x_464 = lean_ctor_get(x_462, 1); -lean_inc(x_464); -lean_dec(x_462); -x_465 = lean_array_push(x_2, x_463); -x_466 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_467 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams), 10, 3); -lean_closure_set(x_467, 0, x_1); -lean_closure_set(x_467, 1, x_465); -lean_closure_set(x_467, 2, x_466); -x_468 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(x_427, x_467, x_4, x_5, x_6, x_7, x_8, x_9, x_464); -return x_468; +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_47 = lean_ctor_get(x_46, 0); +lean_inc(x_47); +x_48 = lean_ctor_get(x_46, 1); +lean_inc(x_48); +lean_dec(x_46); +x_49 = lean_array_push(x_3, x_47); +x_50 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams), 10, 3); +lean_closure_set(x_50, 0, x_20); +lean_closure_set(x_50, 1, x_2); +lean_closure_set(x_50, 2, x_49); +x_51 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(x_15, x_50, x_5, x_6, x_7, x_8, x_9, x_10, x_48); +return x_51; } else { -uint8_t x_469; -lean_dec(x_427); +uint8_t x_52; +lean_dec(x_20); +lean_dec(x_15); +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_469 = !lean_is_exclusive(x_462); -if (x_469 == 0) +x_52 = !lean_is_exclusive(x_46); +if (x_52 == 0) { -return x_462; +return x_46; } else { -lean_object* x_470; lean_object* x_471; lean_object* x_472; -x_470 = lean_ctor_get(x_462, 0); -x_471 = lean_ctor_get(x_462, 1); -lean_inc(x_471); -lean_inc(x_470); -lean_dec(x_462); -x_472 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_472, 0, x_470); -lean_ctor_set(x_472, 1, x_471); -return x_472; +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_46, 0); +x_54 = lean_ctor_get(x_46, 1); +lean_inc(x_54); +lean_inc(x_53); +lean_dec(x_46); +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; } } } case 1: { -lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; -lean_dec(x_457); -x_473 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5; -x_474 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__4___boxed), 11, 3); -lean_closure_set(x_474, 0, x_430); -lean_closure_set(x_474, 1, x_435); -lean_closure_set(x_474, 2, x_473); -x_475 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11; -x_476 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__3___rarg), 9, 2); -lean_closure_set(x_476, 0, x_475); -lean_closure_set(x_476, 1, x_474); +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; +lean_dec(x_41); +x_56 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5; +x_57 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__3___boxed), 11, 3); +lean_closure_set(x_57, 0, x_22); +lean_closure_set(x_57, 1, x_23); +lean_closure_set(x_57, 2, x_56); +x_58 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11; +x_59 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__3___rarg), 9, 2); +lean_closure_set(x_59, 0, x_58); +lean_closure_set(x_59, 1, x_57); +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_477 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingDomain___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__16(x_476, x_4, x_5, x_6, x_7, x_8, x_9, x_439); -if (lean_obj_tag(x_477) == 0) +x_60 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingDomain___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__3(x_59, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +if (lean_obj_tag(x_60) == 0) { -lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; -x_478 = lean_ctor_get(x_477, 0); -lean_inc(x_478); -x_479 = lean_ctor_get(x_477, 1); -lean_inc(x_479); -lean_dec(x_477); -x_480 = lean_array_push(x_2, x_478); -x_481 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_482 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams), 10, 3); -lean_closure_set(x_482, 0, x_1); -lean_closure_set(x_482, 1, x_480); -lean_closure_set(x_482, 2, x_481); -x_483 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(x_427, x_482, x_4, x_5, x_6, x_7, x_8, x_9, x_479); -return x_483; +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_61 = lean_ctor_get(x_60, 0); +lean_inc(x_61); +x_62 = lean_ctor_get(x_60, 1); +lean_inc(x_62); +lean_dec(x_60); +x_63 = lean_array_push(x_3, x_61); +x_64 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams), 10, 3); +lean_closure_set(x_64, 0, x_20); +lean_closure_set(x_64, 1, x_2); +lean_closure_set(x_64, 2, x_63); +x_65 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(x_15, x_64, x_5, x_6, x_7, x_8, x_9, x_10, x_62); +return x_65; } else { -uint8_t x_484; -lean_dec(x_427); +uint8_t x_66; +lean_dec(x_20); +lean_dec(x_15); +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_484 = !lean_is_exclusive(x_477); -if (x_484 == 0) +x_66 = !lean_is_exclusive(x_60); +if (x_66 == 0) { -return x_477; +return x_60; } else { -lean_object* x_485; lean_object* x_486; lean_object* x_487; -x_485 = lean_ctor_get(x_477, 0); -x_486 = lean_ctor_get(x_477, 1); -lean_inc(x_486); -lean_inc(x_485); -lean_dec(x_477); -x_487 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_487, 0, x_485); -lean_ctor_set(x_487, 1, x_486); -return x_487; +lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_60, 0); +x_68 = lean_ctor_get(x_60, 1); +lean_inc(x_68); +lean_inc(x_67); +lean_dec(x_60); +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 2: { -lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; -lean_dec(x_457); -x_488 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5; -x_489 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__5___boxed), 11, 3); -lean_closure_set(x_489, 0, x_430); -lean_closure_set(x_489, 1, x_435); -lean_closure_set(x_489, 2, x_488); -x_490 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11; -x_491 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__3___rarg), 9, 2); -lean_closure_set(x_491, 0, x_490); -lean_closure_set(x_491, 1, x_489); +lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +lean_dec(x_41); +x_70 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5; +x_71 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__4___boxed), 11, 3); +lean_closure_set(x_71, 0, x_22); +lean_closure_set(x_71, 1, x_23); +lean_closure_set(x_71, 2, x_70); +x_72 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11; +x_73 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__3___rarg), 9, 2); +lean_closure_set(x_73, 0, x_72); +lean_closure_set(x_73, 1, x_71); +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_492 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingDomain___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__16(x_491, x_4, x_5, x_6, x_7, x_8, x_9, x_439); -if (lean_obj_tag(x_492) == 0) -{ -lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; -x_493 = lean_ctor_get(x_492, 0); -lean_inc(x_493); -x_494 = lean_ctor_get(x_492, 1); -lean_inc(x_494); -lean_dec(x_492); -x_495 = lean_array_push(x_2, x_493); -x_496 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_497 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams), 10, 3); -lean_closure_set(x_497, 0, x_1); -lean_closure_set(x_497, 1, x_495); -lean_closure_set(x_497, 2, x_496); -x_498 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(x_427, x_497, x_4, x_5, x_6, x_7, x_8, x_9, x_494); -return x_498; +x_74 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingDomain___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__3(x_73, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +if (lean_obj_tag(x_74) == 0) +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_75 = lean_ctor_get(x_74, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_74, 1); +lean_inc(x_76); +lean_dec(x_74); +x_77 = lean_array_push(x_3, x_75); +x_78 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams), 10, 3); +lean_closure_set(x_78, 0, x_20); +lean_closure_set(x_78, 1, x_2); +lean_closure_set(x_78, 2, x_77); +x_79 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(x_15, x_78, x_5, x_6, x_7, x_8, x_9, x_10, x_76); +return x_79; } else { -uint8_t x_499; -lean_dec(x_427); +uint8_t x_80; +lean_dec(x_20); +lean_dec(x_15); +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_499 = !lean_is_exclusive(x_492); -if (x_499 == 0) +x_80 = !lean_is_exclusive(x_74); +if (x_80 == 0) { -return x_492; +return x_74; } else { -lean_object* x_500; lean_object* x_501; lean_object* x_502; -x_500 = lean_ctor_get(x_492, 0); -x_501 = lean_ctor_get(x_492, 1); -lean_inc(x_501); -lean_inc(x_500); -lean_dec(x_492); -x_502 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_502, 0, x_500); -lean_ctor_set(x_502, 1, x_501); -return x_502; +lean_object* x_81; lean_object* x_82; lean_object* x_83; +x_81 = lean_ctor_get(x_74, 0); +x_82 = lean_ctor_get(x_74, 1); +lean_inc(x_82); +lean_inc(x_81); +lean_dec(x_74); +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; } } } case 3: { -lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; -lean_dec(x_457); -x_503 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5; -x_504 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__6___boxed), 11, 3); -lean_closure_set(x_504, 0, x_430); -lean_closure_set(x_504, 1, x_435); -lean_closure_set(x_504, 2, x_503); -x_505 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11; -x_506 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__3___rarg), 9, 2); -lean_closure_set(x_506, 0, x_505); -lean_closure_set(x_506, 1, x_504); +lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; +lean_dec(x_41); +x_84 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5; +x_85 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__5___boxed), 11, 3); +lean_closure_set(x_85, 0, x_22); +lean_closure_set(x_85, 1, x_23); +lean_closure_set(x_85, 2, x_84); +x_86 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11; +x_87 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__3___rarg), 9, 2); +lean_closure_set(x_87, 0, x_86); +lean_closure_set(x_87, 1, x_85); +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_507 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingDomain___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__16(x_506, x_4, x_5, x_6, x_7, x_8, x_9, x_439); -if (lean_obj_tag(x_507) == 0) +x_88 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingDomain___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__3(x_87, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +if (lean_obj_tag(x_88) == 0) { -lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; -x_508 = lean_ctor_get(x_507, 0); -lean_inc(x_508); -x_509 = lean_ctor_get(x_507, 1); -lean_inc(x_509); -lean_dec(x_507); -x_510 = lean_array_push(x_2, x_508); -x_511 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_512 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams), 10, 3); -lean_closure_set(x_512, 0, x_1); -lean_closure_set(x_512, 1, x_510); -lean_closure_set(x_512, 2, x_511); -x_513 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(x_427, x_512, x_4, x_5, x_6, x_7, x_8, x_9, x_509); -return x_513; +lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_89 = lean_ctor_get(x_88, 0); +lean_inc(x_89); +x_90 = lean_ctor_get(x_88, 1); +lean_inc(x_90); +lean_dec(x_88); +x_91 = lean_array_push(x_3, x_89); +x_92 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams), 10, 3); +lean_closure_set(x_92, 0, x_20); +lean_closure_set(x_92, 1, x_2); +lean_closure_set(x_92, 2, x_91); +x_93 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(x_15, x_92, x_5, x_6, x_7, x_8, x_9, x_10, x_90); +return x_93; } else { -uint8_t x_514; -lean_dec(x_427); +uint8_t x_94; +lean_dec(x_20); +lean_dec(x_15); +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_514 = !lean_is_exclusive(x_507); -if (x_514 == 0) +x_94 = !lean_is_exclusive(x_88); +if (x_94 == 0) { -return x_507; +return x_88; } else { -lean_object* x_515; lean_object* x_516; lean_object* x_517; -x_515 = lean_ctor_get(x_507, 0); -x_516 = lean_ctor_get(x_507, 1); -lean_inc(x_516); -lean_inc(x_515); -lean_dec(x_507); -x_517 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_517, 0, x_515); -lean_ctor_set(x_517, 1, x_516); -return x_517; +lean_object* x_95; lean_object* x_96; lean_object* x_97; +x_95 = lean_ctor_get(x_88, 0); +x_96 = lean_ctor_get(x_88, 1); +lean_inc(x_96); +lean_inc(x_95); +lean_dec(x_88); +x_97 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_97, 0, x_95); +lean_ctor_set(x_97, 1, x_96); +return x_97; } } } case 4: { -lean_object* x_518; lean_object* x_519; lean_object* x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; -x_518 = lean_ctor_get(x_457, 0); -lean_inc(x_518); -lean_dec(x_457); -x_519 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11; -x_520 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5; -x_521 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__8___boxed), 13, 5); -lean_closure_set(x_521, 0, x_518); -lean_closure_set(x_521, 1, x_519); -lean_closure_set(x_521, 2, x_430); -lean_closure_set(x_521, 3, x_435); -lean_closure_set(x_521, 4, x_520); -x_522 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__12; -x_523 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__3___rarg), 9, 2); -lean_closure_set(x_523, 0, x_522); -lean_closure_set(x_523, 1, x_521); +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_98 = lean_ctor_get(x_41, 0); +lean_inc(x_98); +lean_dec(x_41); +x_99 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5; +x_100 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__7___boxed), 12, 4); +lean_closure_set(x_100, 0, x_98); +lean_closure_set(x_100, 1, x_22); +lean_closure_set(x_100, 2, x_23); +lean_closure_set(x_100, 3, x_99); +x_101 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___closed__3; +x_102 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__3___rarg), 9, 2); +lean_closure_set(x_102, 0, x_101); +lean_closure_set(x_102, 1, x_100); +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_524 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingDomain___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__16(x_523, x_4, x_5, x_6, x_7, x_8, x_9, x_439); -if (lean_obj_tag(x_524) == 0) +x_103 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingDomain___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__3(x_102, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +if (lean_obj_tag(x_103) == 0) { -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_525 = lean_ctor_get(x_524, 0); -lean_inc(x_525); -x_526 = lean_ctor_get(x_524, 1); -lean_inc(x_526); -lean_dec(x_524); -x_527 = lean_array_push(x_2, x_525); -x_528 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_529 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams), 10, 3); -lean_closure_set(x_529, 0, x_1); -lean_closure_set(x_529, 1, x_527); -lean_closure_set(x_529, 2, x_528); -x_530 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(x_427, x_529, x_4, x_5, x_6, x_7, x_8, x_9, x_526); -return x_530; +lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; +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_array_push(x_3, x_104); +x_107 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams), 10, 3); +lean_closure_set(x_107, 0, x_20); +lean_closure_set(x_107, 1, x_2); +lean_closure_set(x_107, 2, x_106); +x_108 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(x_15, x_107, x_5, x_6, x_7, x_8, x_9, x_10, x_105); +return x_108; } else { -uint8_t x_531; -lean_dec(x_427); +uint8_t x_109; +lean_dec(x_20); +lean_dec(x_15); +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_531 = !lean_is_exclusive(x_524); -if (x_531 == 0) +x_109 = !lean_is_exclusive(x_103); +if (x_109 == 0) { -return x_524; +return x_103; } else { -lean_object* x_532; lean_object* x_533; lean_object* x_534; -x_532 = lean_ctor_get(x_524, 0); -x_533 = lean_ctor_get(x_524, 1); -lean_inc(x_533); -lean_inc(x_532); -lean_dec(x_524); -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; +lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_110 = lean_ctor_get(x_103, 0); +x_111 = lean_ctor_get(x_103, 1); +lean_inc(x_111); +lean_inc(x_110); +lean_dec(x_103); +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; } } } case 5: { -lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; -lean_dec(x_457); -x_535 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5; -x_536 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__9___boxed), 11, 3); -lean_closure_set(x_536, 0, x_430); -lean_closure_set(x_536, 1, x_435); -lean_closure_set(x_536, 2, x_535); -x_537 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11; -x_538 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__3___rarg), 9, 2); -lean_closure_set(x_538, 0, x_537); -lean_closure_set(x_538, 1, x_536); +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; +lean_dec(x_41); +x_113 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5; +x_114 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__8___boxed), 11, 3); +lean_closure_set(x_114, 0, x_22); +lean_closure_set(x_114, 1, x_23); +lean_closure_set(x_114, 2, x_113); +x_115 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11; +x_116 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__3___rarg), 9, 2); +lean_closure_set(x_116, 0, x_115); +lean_closure_set(x_116, 1, x_114); +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_539 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingDomain___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__16(x_538, x_4, x_5, x_6, x_7, x_8, x_9, x_439); -if (lean_obj_tag(x_539) == 0) +x_117 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingDomain___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__3(x_116, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +if (lean_obj_tag(x_117) == 0) { -lean_object* x_540; lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_545; -x_540 = lean_ctor_get(x_539, 0); -lean_inc(x_540); -x_541 = lean_ctor_get(x_539, 1); -lean_inc(x_541); -lean_dec(x_539); -x_542 = lean_array_push(x_2, x_540); -x_543 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_544 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams), 10, 3); -lean_closure_set(x_544, 0, x_1); -lean_closure_set(x_544, 1, x_542); -lean_closure_set(x_544, 2, x_543); -x_545 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(x_427, x_544, x_4, x_5, x_6, x_7, x_8, x_9, x_541); -return x_545; +lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; +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_array_push(x_3, x_118); +x_121 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams), 10, 3); +lean_closure_set(x_121, 0, x_20); +lean_closure_set(x_121, 1, x_2); +lean_closure_set(x_121, 2, x_120); +x_122 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(x_15, x_121, x_5, x_6, x_7, x_8, x_9, x_10, x_119); +return x_122; } else { -uint8_t x_546; -lean_dec(x_427); +uint8_t x_123; +lean_dec(x_20); +lean_dec(x_15); +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_546 = !lean_is_exclusive(x_539); -if (x_546 == 0) +x_123 = !lean_is_exclusive(x_117); +if (x_123 == 0) { -return x_539; +return x_117; } else { -lean_object* x_547; lean_object* x_548; lean_object* x_549; -x_547 = lean_ctor_get(x_539, 0); -x_548 = lean_ctor_get(x_539, 1); -lean_inc(x_548); -lean_inc(x_547); -lean_dec(x_539); -x_549 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_549, 0, x_547); -lean_ctor_set(x_549, 1, x_548); -return x_549; +lean_object* x_124; lean_object* x_125; lean_object* x_126; +x_124 = lean_ctor_get(x_117, 0); +x_125 = lean_ctor_get(x_117, 1); +lean_inc(x_125); +lean_inc(x_124); +lean_dec(x_117); +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; } } } case 6: { -lean_object* x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; lean_object* x_554; -lean_dec(x_457); -x_550 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5; -x_551 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__10___boxed), 11, 3); -lean_closure_set(x_551, 0, x_430); -lean_closure_set(x_551, 1, x_435); -lean_closure_set(x_551, 2, x_550); -x_552 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11; -x_553 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__3___rarg), 9, 2); -lean_closure_set(x_553, 0, x_552); -lean_closure_set(x_553, 1, x_551); +lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; +lean_dec(x_41); +x_127 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5; +x_128 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__9___boxed), 11, 3); +lean_closure_set(x_128, 0, x_22); +lean_closure_set(x_128, 1, x_23); +lean_closure_set(x_128, 2, x_127); +x_129 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11; +x_130 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__3___rarg), 9, 2); +lean_closure_set(x_130, 0, x_129); +lean_closure_set(x_130, 1, x_128); +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_554 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingDomain___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__16(x_553, x_4, x_5, x_6, x_7, x_8, x_9, x_439); -if (lean_obj_tag(x_554) == 0) +x_131 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingDomain___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__3(x_130, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +if (lean_obj_tag(x_131) == 0) { -lean_object* x_555; lean_object* x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; lean_object* x_560; -x_555 = lean_ctor_get(x_554, 0); -lean_inc(x_555); -x_556 = lean_ctor_get(x_554, 1); -lean_inc(x_556); -lean_dec(x_554); -x_557 = lean_array_push(x_2, x_555); -x_558 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_559 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams), 10, 3); -lean_closure_set(x_559, 0, x_1); -lean_closure_set(x_559, 1, x_557); -lean_closure_set(x_559, 2, x_558); -x_560 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(x_427, x_559, x_4, x_5, x_6, x_7, x_8, x_9, x_556); -return x_560; +lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; +x_132 = lean_ctor_get(x_131, 0); +lean_inc(x_132); +x_133 = lean_ctor_get(x_131, 1); +lean_inc(x_133); +lean_dec(x_131); +x_134 = lean_array_push(x_3, x_132); +x_135 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams), 10, 3); +lean_closure_set(x_135, 0, x_20); +lean_closure_set(x_135, 1, x_2); +lean_closure_set(x_135, 2, x_134); +x_136 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(x_15, x_135, x_5, x_6, x_7, x_8, x_9, x_10, x_133); +return x_136; } else { -uint8_t x_561; -lean_dec(x_427); +uint8_t x_137; +lean_dec(x_20); +lean_dec(x_15); +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_561 = !lean_is_exclusive(x_554); -if (x_561 == 0) +x_137 = !lean_is_exclusive(x_131); +if (x_137 == 0) { -return x_554; +return x_131; } else { -lean_object* x_562; lean_object* x_563; lean_object* x_564; -x_562 = lean_ctor_get(x_554, 0); -x_563 = lean_ctor_get(x_554, 1); -lean_inc(x_563); -lean_inc(x_562); -lean_dec(x_554); -x_564 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_564, 0, x_562); -lean_ctor_set(x_564, 1, x_563); -return x_564; +lean_object* x_138; lean_object* x_139; lean_object* x_140; +x_138 = lean_ctor_get(x_131, 0); +x_139 = lean_ctor_get(x_131, 1); +lean_inc(x_139); +lean_inc(x_138); +lean_dec(x_131); +x_140 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_140, 0, x_138); +lean_ctor_set(x_140, 1, x_139); +return x_140; } } } case 7: { -lean_object* x_565; lean_object* x_566; lean_object* x_567; lean_object* x_568; lean_object* x_569; -lean_dec(x_457); -x_565 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5; -x_566 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__11___boxed), 11, 3); -lean_closure_set(x_566, 0, x_430); -lean_closure_set(x_566, 1, x_435); -lean_closure_set(x_566, 2, x_565); -x_567 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11; -x_568 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__3___rarg), 9, 2); -lean_closure_set(x_568, 0, x_567); -lean_closure_set(x_568, 1, x_566); +lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; +lean_dec(x_41); +x_141 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5; +x_142 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__10___boxed), 11, 3); +lean_closure_set(x_142, 0, x_22); +lean_closure_set(x_142, 1, x_23); +lean_closure_set(x_142, 2, x_141); +x_143 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11; +x_144 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__3___rarg), 9, 2); +lean_closure_set(x_144, 0, x_143); +lean_closure_set(x_144, 1, x_142); +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_569 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingDomain___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__16(x_568, x_4, x_5, x_6, x_7, x_8, x_9, x_439); -if (lean_obj_tag(x_569) == 0) +x_145 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingDomain___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__3(x_144, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +if (lean_obj_tag(x_145) == 0) { -lean_object* x_570; lean_object* x_571; lean_object* x_572; lean_object* x_573; lean_object* x_574; lean_object* x_575; -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_array_push(x_2, x_570); -x_573 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_574 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams), 10, 3); -lean_closure_set(x_574, 0, x_1); -lean_closure_set(x_574, 1, x_572); -lean_closure_set(x_574, 2, x_573); -x_575 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(x_427, x_574, x_4, x_5, x_6, x_7, x_8, x_9, x_571); -return x_575; +lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; +x_146 = lean_ctor_get(x_145, 0); +lean_inc(x_146); +x_147 = lean_ctor_get(x_145, 1); +lean_inc(x_147); +lean_dec(x_145); +x_148 = lean_array_push(x_3, x_146); +x_149 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams), 10, 3); +lean_closure_set(x_149, 0, x_20); +lean_closure_set(x_149, 1, x_2); +lean_closure_set(x_149, 2, x_148); +x_150 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(x_15, x_149, x_5, x_6, x_7, x_8, x_9, x_10, x_147); +return x_150; } else { -uint8_t x_576; -lean_dec(x_427); +uint8_t x_151; +lean_dec(x_20); +lean_dec(x_15); +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_576 = !lean_is_exclusive(x_569); -if (x_576 == 0) +x_151 = !lean_is_exclusive(x_145); +if (x_151 == 0) { -return x_569; +return x_145; } else { -lean_object* x_577; lean_object* x_578; lean_object* x_579; -x_577 = lean_ctor_get(x_569, 0); -x_578 = lean_ctor_get(x_569, 1); -lean_inc(x_578); -lean_inc(x_577); -lean_dec(x_569); -x_579 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_579, 0, x_577); -lean_ctor_set(x_579, 1, x_578); -return x_579; +lean_object* x_152; lean_object* x_153; lean_object* x_154; +x_152 = lean_ctor_get(x_145, 0); +x_153 = lean_ctor_get(x_145, 1); +lean_inc(x_153); +lean_inc(x_152); +lean_dec(x_145); +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; } } } case 8: { -lean_object* x_580; lean_object* x_581; lean_object* x_582; lean_object* x_583; lean_object* x_584; -lean_dec(x_457); -x_580 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5; -x_581 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__12___boxed), 11, 3); -lean_closure_set(x_581, 0, x_430); -lean_closure_set(x_581, 1, x_435); -lean_closure_set(x_581, 2, x_580); -x_582 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11; -x_583 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__3___rarg), 9, 2); -lean_closure_set(x_583, 0, x_582); -lean_closure_set(x_583, 1, x_581); +lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; +lean_dec(x_41); +x_155 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5; +x_156 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__11___boxed), 11, 3); +lean_closure_set(x_156, 0, x_22); +lean_closure_set(x_156, 1, x_23); +lean_closure_set(x_156, 2, x_155); +x_157 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11; +x_158 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__3___rarg), 9, 2); +lean_closure_set(x_158, 0, x_157); +lean_closure_set(x_158, 1, x_156); +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_584 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingDomain___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__16(x_583, x_4, x_5, x_6, x_7, x_8, x_9, x_439); -if (lean_obj_tag(x_584) == 0) +x_159 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingDomain___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__3(x_158, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +if (lean_obj_tag(x_159) == 0) { -lean_object* x_585; lean_object* x_586; lean_object* x_587; lean_object* x_588; lean_object* x_589; lean_object* x_590; -x_585 = lean_ctor_get(x_584, 0); -lean_inc(x_585); -x_586 = lean_ctor_get(x_584, 1); -lean_inc(x_586); -lean_dec(x_584); -x_587 = lean_array_push(x_2, x_585); -x_588 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_589 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams), 10, 3); -lean_closure_set(x_589, 0, x_1); -lean_closure_set(x_589, 1, x_587); -lean_closure_set(x_589, 2, x_588); -x_590 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(x_427, x_589, x_4, x_5, x_6, x_7, x_8, x_9, x_586); -return x_590; +lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; +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 = lean_array_push(x_3, x_160); +x_163 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams), 10, 3); +lean_closure_set(x_163, 0, x_20); +lean_closure_set(x_163, 1, x_2); +lean_closure_set(x_163, 2, x_162); +x_164 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(x_15, x_163, x_5, x_6, x_7, x_8, x_9, x_10, x_161); +return x_164; } else { -uint8_t x_591; -lean_dec(x_427); +uint8_t x_165; +lean_dec(x_20); +lean_dec(x_15); +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_591 = !lean_is_exclusive(x_584); -if (x_591 == 0) +x_165 = !lean_is_exclusive(x_159); +if (x_165 == 0) { -return x_584; +return x_159; } else { -lean_object* x_592; lean_object* x_593; lean_object* x_594; -x_592 = lean_ctor_get(x_584, 0); -x_593 = lean_ctor_get(x_584, 1); -lean_inc(x_593); -lean_inc(x_592); -lean_dec(x_584); -x_594 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_594, 0, x_592); -lean_ctor_set(x_594, 1, x_593); -return x_594; +lean_object* x_166; lean_object* x_167; lean_object* x_168; +x_166 = lean_ctor_get(x_159, 0); +x_167 = lean_ctor_get(x_159, 1); +lean_inc(x_167); +lean_inc(x_166); +lean_dec(x_159); +x_168 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_168, 0, x_166); +lean_ctor_set(x_168, 1, x_167); +return x_168; } } } case 9: { -lean_object* x_595; lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; -lean_dec(x_457); -x_595 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5; -x_596 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__13___boxed), 11, 3); -lean_closure_set(x_596, 0, x_430); -lean_closure_set(x_596, 1, x_435); -lean_closure_set(x_596, 2, x_595); -x_597 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11; -x_598 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__3___rarg), 9, 2); -lean_closure_set(x_598, 0, x_597); -lean_closure_set(x_598, 1, x_596); +lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; +lean_dec(x_41); +x_169 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5; +x_170 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__12___boxed), 11, 3); +lean_closure_set(x_170, 0, x_22); +lean_closure_set(x_170, 1, x_23); +lean_closure_set(x_170, 2, x_169); +x_171 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11; +x_172 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__3___rarg), 9, 2); +lean_closure_set(x_172, 0, x_171); +lean_closure_set(x_172, 1, x_170); +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_599 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingDomain___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__16(x_598, x_4, x_5, x_6, x_7, x_8, x_9, x_439); -if (lean_obj_tag(x_599) == 0) +x_173 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingDomain___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__3(x_172, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +if (lean_obj_tag(x_173) == 0) { -lean_object* x_600; lean_object* x_601; lean_object* x_602; lean_object* x_603; lean_object* x_604; lean_object* x_605; -x_600 = lean_ctor_get(x_599, 0); -lean_inc(x_600); -x_601 = lean_ctor_get(x_599, 1); -lean_inc(x_601); -lean_dec(x_599); -x_602 = lean_array_push(x_2, x_600); -x_603 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_604 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams), 10, 3); -lean_closure_set(x_604, 0, x_1); -lean_closure_set(x_604, 1, x_602); -lean_closure_set(x_604, 2, x_603); -x_605 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(x_427, x_604, x_4, x_5, x_6, x_7, x_8, x_9, x_601); -return x_605; +lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; +x_174 = lean_ctor_get(x_173, 0); +lean_inc(x_174); +x_175 = lean_ctor_get(x_173, 1); +lean_inc(x_175); +lean_dec(x_173); +x_176 = lean_array_push(x_3, x_174); +x_177 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams), 10, 3); +lean_closure_set(x_177, 0, x_20); +lean_closure_set(x_177, 1, x_2); +lean_closure_set(x_177, 2, x_176); +x_178 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(x_15, x_177, x_5, x_6, x_7, x_8, x_9, x_10, x_175); +return x_178; } else { -uint8_t x_606; -lean_dec(x_427); +uint8_t x_179; +lean_dec(x_20); +lean_dec(x_15); +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_606 = !lean_is_exclusive(x_599); -if (x_606 == 0) +x_179 = !lean_is_exclusive(x_173); +if (x_179 == 0) { -return x_599; +return x_173; } else { -lean_object* x_607; lean_object* x_608; lean_object* x_609; -x_607 = lean_ctor_get(x_599, 0); -x_608 = lean_ctor_get(x_599, 1); -lean_inc(x_608); -lean_inc(x_607); -lean_dec(x_599); -x_609 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_609, 0, x_607); -lean_ctor_set(x_609, 1, x_608); -return x_609; +lean_object* x_180; lean_object* x_181; lean_object* x_182; +x_180 = lean_ctor_get(x_173, 0); +x_181 = lean_ctor_get(x_173, 1); +lean_inc(x_181); +lean_inc(x_180); +lean_dec(x_173); +x_182 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_182, 0, x_180); +lean_ctor_set(x_182, 1, x_181); +return x_182; } } } case 10: { -lean_object* x_610; lean_object* x_611; lean_object* x_612; lean_object* x_613; lean_object* x_614; -lean_dec(x_457); -x_610 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5; -x_611 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__14___boxed), 11, 3); -lean_closure_set(x_611, 0, x_430); -lean_closure_set(x_611, 1, x_435); -lean_closure_set(x_611, 2, x_610); -x_612 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11; -x_613 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__3___rarg), 9, 2); -lean_closure_set(x_613, 0, x_612); -lean_closure_set(x_613, 1, x_611); +lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; +lean_dec(x_41); +x_183 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5; +x_184 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__13___boxed), 11, 3); +lean_closure_set(x_184, 0, x_22); +lean_closure_set(x_184, 1, x_23); +lean_closure_set(x_184, 2, x_183); +x_185 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11; +x_186 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__3___rarg), 9, 2); +lean_closure_set(x_186, 0, x_185); +lean_closure_set(x_186, 1, x_184); +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_614 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingDomain___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__16(x_613, x_4, x_5, x_6, x_7, x_8, x_9, x_439); -if (lean_obj_tag(x_614) == 0) +x_187 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingDomain___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__3(x_186, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +if (lean_obj_tag(x_187) == 0) { -lean_object* x_615; lean_object* x_616; lean_object* x_617; lean_object* x_618; lean_object* x_619; lean_object* x_620; -x_615 = lean_ctor_get(x_614, 0); -lean_inc(x_615); -x_616 = lean_ctor_get(x_614, 1); -lean_inc(x_616); -lean_dec(x_614); -x_617 = lean_array_push(x_2, x_615); -x_618 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_619 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams), 10, 3); -lean_closure_set(x_619, 0, x_1); -lean_closure_set(x_619, 1, x_617); -lean_closure_set(x_619, 2, x_618); -x_620 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(x_427, x_619, x_4, x_5, x_6, x_7, x_8, x_9, x_616); -return x_620; +lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; +x_188 = lean_ctor_get(x_187, 0); +lean_inc(x_188); +x_189 = lean_ctor_get(x_187, 1); +lean_inc(x_189); +lean_dec(x_187); +x_190 = lean_array_push(x_3, x_188); +x_191 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams), 10, 3); +lean_closure_set(x_191, 0, x_20); +lean_closure_set(x_191, 1, x_2); +lean_closure_set(x_191, 2, x_190); +x_192 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(x_15, x_191, x_5, x_6, x_7, x_8, x_9, x_10, x_189); +return x_192; } else { -uint8_t x_621; -lean_dec(x_427); +uint8_t x_193; +lean_dec(x_20); +lean_dec(x_15); +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_621 = !lean_is_exclusive(x_614); -if (x_621 == 0) +x_193 = !lean_is_exclusive(x_187); +if (x_193 == 0) { -return x_614; +return x_187; } else { -lean_object* x_622; lean_object* x_623; lean_object* x_624; -x_622 = lean_ctor_get(x_614, 0); -x_623 = lean_ctor_get(x_614, 1); -lean_inc(x_623); -lean_inc(x_622); -lean_dec(x_614); -x_624 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_624, 0, x_622); -lean_ctor_set(x_624, 1, x_623); -return x_624; +lean_object* x_194; lean_object* x_195; lean_object* x_196; +x_194 = lean_ctor_get(x_187, 0); +x_195 = lean_ctor_get(x_187, 1); +lean_inc(x_195); +lean_inc(x_194); +lean_dec(x_187); +x_196 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_196, 0, x_194); +lean_ctor_set(x_196, 1, x_195); +return x_196; } } } default: { -lean_object* x_625; lean_object* x_626; lean_object* x_627; lean_object* x_628; lean_object* x_629; -lean_dec(x_457); -x_625 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5; -x_626 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__15___boxed), 11, 3); -lean_closure_set(x_626, 0, x_430); -lean_closure_set(x_626, 1, x_435); -lean_closure_set(x_626, 2, x_625); -x_627 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11; -x_628 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__3___rarg), 9, 2); -lean_closure_set(x_628, 0, x_627); -lean_closure_set(x_628, 1, x_626); +lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; +lean_dec(x_41); +x_197 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5; +x_198 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__14___boxed), 11, 3); +lean_closure_set(x_198, 0, x_22); +lean_closure_set(x_198, 1, x_23); +lean_closure_set(x_198, 2, x_197); +x_199 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11; +x_200 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__3___rarg), 9, 2); +lean_closure_set(x_200, 0, x_199); +lean_closure_set(x_200, 1, x_198); +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_629 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingDomain___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__16(x_628, x_4, x_5, x_6, x_7, x_8, x_9, x_439); -if (lean_obj_tag(x_629) == 0) +x_201 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingDomain___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__3(x_200, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +if (lean_obj_tag(x_201) == 0) { -lean_object* x_630; lean_object* x_631; lean_object* x_632; lean_object* x_633; lean_object* x_634; lean_object* x_635; -x_630 = lean_ctor_get(x_629, 0); -lean_inc(x_630); -x_631 = lean_ctor_get(x_629, 1); -lean_inc(x_631); -lean_dec(x_629); -x_632 = lean_array_push(x_2, x_630); -x_633 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_634 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams), 10, 3); -lean_closure_set(x_634, 0, x_1); -lean_closure_set(x_634, 1, x_632); -lean_closure_set(x_634, 2, x_633); -x_635 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(x_427, x_634, x_4, x_5, x_6, x_7, x_8, x_9, x_631); -return x_635; +lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; +x_202 = lean_ctor_get(x_201, 0); +lean_inc(x_202); +x_203 = lean_ctor_get(x_201, 1); +lean_inc(x_203); +lean_dec(x_201); +x_204 = lean_array_push(x_3, x_202); +x_205 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams), 10, 3); +lean_closure_set(x_205, 0, x_20); +lean_closure_set(x_205, 1, x_2); +lean_closure_set(x_205, 2, x_204); +x_206 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(x_15, x_205, x_5, x_6, x_7, x_8, x_9, x_10, x_203); +return x_206; } else { -uint8_t x_636; -lean_dec(x_427); +uint8_t x_207; +lean_dec(x_20); +lean_dec(x_15); +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_636 = !lean_is_exclusive(x_629); -if (x_636 == 0) +x_207 = !lean_is_exclusive(x_201); +if (x_207 == 0) { -return x_629; +return x_201; } else { -lean_object* x_637; lean_object* x_638; lean_object* x_639; -x_637 = lean_ctor_get(x_629, 0); -x_638 = lean_ctor_get(x_629, 1); -lean_inc(x_638); -lean_inc(x_637); -lean_dec(x_629); -x_639 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_639, 0, x_637); -lean_ctor_set(x_639, 1, x_638); -return x_639; +lean_object* x_208; lean_object* x_209; lean_object* x_210; +x_208 = lean_ctor_get(x_201, 0); +x_209 = lean_ctor_get(x_201, 1); +lean_inc(x_209); +lean_inc(x_208); +lean_dec(x_201); +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; } } } @@ -43951,286 +42950,278 @@ return x_639; } else { -lean_object* x_640; lean_object* x_641; lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_645; -lean_dec(x_428); -x_640 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11; -x_641 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5; -x_642 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__16___boxed), 12, 4); -lean_closure_set(x_642, 0, x_640); -lean_closure_set(x_642, 1, x_430); -lean_closure_set(x_642, 2, x_435); -lean_closure_set(x_642, 3, x_641); -x_643 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__14; -x_644 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__3___rarg), 9, 2); -lean_closure_set(x_644, 0, x_643); -lean_closure_set(x_644, 1, x_642); +lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; +lean_dec(x_16); +x_211 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5; +x_212 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__15___boxed), 11, 3); +lean_closure_set(x_212, 0, x_22); +lean_closure_set(x_212, 1, x_23); +lean_closure_set(x_212, 2, x_211); +x_213 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___closed__4; +x_214 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__3___rarg), 9, 2); +lean_closure_set(x_214, 0, x_213); +lean_closure_set(x_214, 1, x_212); +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_645 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingDomain___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__16(x_644, x_4, x_5, x_6, x_7, x_8, x_9, x_439); -if (lean_obj_tag(x_645) == 0) +x_215 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingDomain___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__3(x_214, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +if (lean_obj_tag(x_215) == 0) { -lean_object* x_646; lean_object* x_647; lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; -x_646 = lean_ctor_get(x_645, 0); -lean_inc(x_646); -x_647 = lean_ctor_get(x_645, 1); -lean_inc(x_647); -lean_dec(x_645); -x_648 = lean_array_push(x_2, x_646); -x_649 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_650 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams), 10, 3); -lean_closure_set(x_650, 0, x_1); -lean_closure_set(x_650, 1, x_648); -lean_closure_set(x_650, 2, x_649); -x_651 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(x_427, x_650, x_4, x_5, x_6, x_7, x_8, x_9, x_647); -return x_651; +lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; +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 = lean_array_push(x_3, x_216); +x_219 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams), 10, 3); +lean_closure_set(x_219, 0, x_20); +lean_closure_set(x_219, 1, x_2); +lean_closure_set(x_219, 2, x_218); +x_220 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(x_15, x_219, x_5, x_6, x_7, x_8, x_9, x_10, x_217); +return x_220; } else { -uint8_t x_652; -lean_dec(x_427); +uint8_t x_221; +lean_dec(x_20); +lean_dec(x_15); +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_652 = !lean_is_exclusive(x_645); -if (x_652 == 0) +x_221 = !lean_is_exclusive(x_215); +if (x_221 == 0) { -return x_645; +return x_215; } else { -lean_object* x_653; lean_object* x_654; lean_object* x_655; -x_653 = lean_ctor_get(x_645, 0); -x_654 = lean_ctor_get(x_645, 1); -lean_inc(x_654); -lean_inc(x_653); -lean_dec(x_645); -x_655 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_655, 0, x_653); -lean_ctor_set(x_655, 1, x_654); -return x_655; +lean_object* x_222; lean_object* x_223; lean_object* x_224; +x_222 = lean_ctor_get(x_215, 0); +x_223 = lean_ctor_get(x_215, 1); +lean_inc(x_223); +lean_inc(x_222); +lean_dec(x_215); +x_224 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_224, 0, x_222); +lean_ctor_set(x_224, 1, x_223); +return x_224; } } } } case 1: { -lean_object* x_656; lean_object* x_657; lean_object* x_658; lean_object* x_659; lean_object* x_660; lean_object* x_661; -lean_dec(x_428); -x_656 = lean_ctor_get(x_436, 1); -lean_inc(x_656); -lean_dec(x_436); -x_657 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5; -x_658 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__17___boxed), 11, 3); -lean_closure_set(x_658, 0, x_430); -lean_closure_set(x_658, 1, x_435); -lean_closure_set(x_658, 2, x_657); -x_659 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11; -x_660 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__3___rarg), 9, 2); -lean_closure_set(x_660, 0, x_659); -lean_closure_set(x_660, 1, x_658); +lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; +lean_dec(x_21); +lean_dec(x_16); +x_225 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5; +x_226 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__16___boxed), 11, 3); +lean_closure_set(x_226, 0, x_22); +lean_closure_set(x_226, 1, x_23); +lean_closure_set(x_226, 2, x_225); +x_227 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11; +x_228 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__3___rarg), 9, 2); +lean_closure_set(x_228, 0, x_227); +lean_closure_set(x_228, 1, x_226); +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_661 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingDomain___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__16(x_660, x_4, x_5, x_6, x_7, x_8, x_9, x_656); -if (lean_obj_tag(x_661) == 0) +x_229 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingDomain___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__3(x_228, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +if (lean_obj_tag(x_229) == 0) { -lean_object* x_662; lean_object* x_663; lean_object* x_664; lean_object* x_665; lean_object* x_666; lean_object* x_667; -x_662 = lean_ctor_get(x_661, 0); -lean_inc(x_662); -x_663 = lean_ctor_get(x_661, 1); -lean_inc(x_663); -lean_dec(x_661); -x_664 = lean_array_push(x_2, x_662); -x_665 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_666 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams), 10, 3); -lean_closure_set(x_666, 0, x_1); -lean_closure_set(x_666, 1, x_664); -lean_closure_set(x_666, 2, x_665); -x_667 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(x_427, x_666, x_4, x_5, x_6, x_7, x_8, x_9, x_663); -return x_667; +lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; +x_230 = lean_ctor_get(x_229, 0); +lean_inc(x_230); +x_231 = lean_ctor_get(x_229, 1); +lean_inc(x_231); +lean_dec(x_229); +x_232 = lean_array_push(x_3, x_230); +x_233 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams), 10, 3); +lean_closure_set(x_233, 0, x_20); +lean_closure_set(x_233, 1, x_2); +lean_closure_set(x_233, 2, x_232); +x_234 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(x_15, x_233, x_5, x_6, x_7, x_8, x_9, x_10, x_231); +return x_234; } else { -uint8_t x_668; -lean_dec(x_427); +uint8_t x_235; +lean_dec(x_20); +lean_dec(x_15); +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_668 = !lean_is_exclusive(x_661); -if (x_668 == 0) +x_235 = !lean_is_exclusive(x_229); +if (x_235 == 0) { -return x_661; +return x_229; } else { -lean_object* x_669; lean_object* x_670; lean_object* x_671; -x_669 = lean_ctor_get(x_661, 0); -x_670 = lean_ctor_get(x_661, 1); -lean_inc(x_670); -lean_inc(x_669); -lean_dec(x_661); -x_671 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_671, 0, x_669); -lean_ctor_set(x_671, 1, x_670); -return x_671; +lean_object* x_236; lean_object* x_237; lean_object* x_238; +x_236 = lean_ctor_get(x_229, 0); +x_237 = lean_ctor_get(x_229, 1); +lean_inc(x_237); +lean_inc(x_236); +lean_dec(x_229); +x_238 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_238, 0, x_236); +lean_ctor_set(x_238, 1, x_237); +return x_238; } } } case 2: { -lean_object* x_672; lean_object* x_673; lean_object* x_674; lean_object* x_675; lean_object* x_676; lean_object* x_677; -lean_dec(x_428); -x_672 = lean_ctor_get(x_436, 1); -lean_inc(x_672); -lean_dec(x_436); -x_673 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5; -x_674 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__18___boxed), 11, 3); -lean_closure_set(x_674, 0, x_430); -lean_closure_set(x_674, 1, x_435); -lean_closure_set(x_674, 2, x_673); -x_675 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11; -x_676 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__3___rarg), 9, 2); -lean_closure_set(x_676, 0, x_675); -lean_closure_set(x_676, 1, x_674); +lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; +lean_dec(x_21); +lean_dec(x_16); +x_239 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5; +x_240 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__17___boxed), 11, 3); +lean_closure_set(x_240, 0, x_22); +lean_closure_set(x_240, 1, x_23); +lean_closure_set(x_240, 2, x_239); +x_241 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11; +x_242 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__3___rarg), 9, 2); +lean_closure_set(x_242, 0, x_241); +lean_closure_set(x_242, 1, x_240); +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_677 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingDomain___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__16(x_676, x_4, x_5, x_6, x_7, x_8, x_9, x_672); -if (lean_obj_tag(x_677) == 0) +x_243 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingDomain___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__3(x_242, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +if (lean_obj_tag(x_243) == 0) { -lean_object* x_678; lean_object* x_679; lean_object* x_680; lean_object* x_681; lean_object* x_682; lean_object* x_683; -x_678 = lean_ctor_get(x_677, 0); -lean_inc(x_678); -x_679 = lean_ctor_get(x_677, 1); -lean_inc(x_679); -lean_dec(x_677); -x_680 = lean_array_push(x_2, x_678); -x_681 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_682 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams), 10, 3); -lean_closure_set(x_682, 0, x_1); -lean_closure_set(x_682, 1, x_680); -lean_closure_set(x_682, 2, x_681); -x_683 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(x_427, x_682, x_4, x_5, x_6, x_7, x_8, x_9, x_679); -return x_683; +lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; +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 = lean_array_push(x_3, x_244); +x_247 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams), 10, 3); +lean_closure_set(x_247, 0, x_20); +lean_closure_set(x_247, 1, x_2); +lean_closure_set(x_247, 2, x_246); +x_248 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(x_15, x_247, x_5, x_6, x_7, x_8, x_9, x_10, x_245); +return x_248; } else { -uint8_t x_684; -lean_dec(x_427); +uint8_t x_249; +lean_dec(x_20); +lean_dec(x_15); +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_684 = !lean_is_exclusive(x_677); -if (x_684 == 0) +x_249 = !lean_is_exclusive(x_243); +if (x_249 == 0) { -return x_677; +return x_243; } else { -lean_object* x_685; lean_object* x_686; lean_object* x_687; -x_685 = lean_ctor_get(x_677, 0); -x_686 = lean_ctor_get(x_677, 1); -lean_inc(x_686); -lean_inc(x_685); -lean_dec(x_677); -x_687 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_687, 0, x_685); -lean_ctor_set(x_687, 1, x_686); -return x_687; +lean_object* x_250; lean_object* x_251; lean_object* x_252; +x_250 = lean_ctor_get(x_243, 0); +x_251 = lean_ctor_get(x_243, 1); +lean_inc(x_251); +lean_inc(x_250); +lean_dec(x_243); +x_252 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_252, 0, x_250); +lean_ctor_set(x_252, 1, x_251); +return x_252; } } } default: { -lean_object* x_688; lean_object* x_689; lean_object* x_690; lean_object* x_691; lean_object* x_692; -lean_dec(x_428); -x_688 = lean_ctor_get(x_436, 1); -lean_inc(x_688); -lean_dec(x_436); -x_689 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__19___boxed), 9, 1); -lean_closure_set(x_689, 0, x_435); -x_690 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11; -x_691 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__3___rarg), 9, 2); -lean_closure_set(x_691, 0, x_690); -lean_closure_set(x_691, 1, x_689); +lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; +lean_dec(x_23); +lean_dec(x_16); +x_253 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__18___boxed), 9, 1); +lean_closure_set(x_253, 0, x_21); +x_254 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11; +x_255 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__3___rarg), 9, 2); +lean_closure_set(x_255, 0, x_254); +lean_closure_set(x_255, 1, x_253); +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_692 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingDomain___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__16(x_691, x_4, x_5, x_6, x_7, x_8, x_9, x_688); -if (lean_obj_tag(x_692) == 0) -{ -lean_object* x_693; lean_object* x_694; lean_object* x_695; lean_object* x_696; lean_object* x_697; lean_object* x_698; -x_693 = lean_ctor_get(x_692, 0); -lean_inc(x_693); -x_694 = lean_ctor_get(x_692, 1); -lean_inc(x_694); -lean_dec(x_692); -x_695 = lean_array_push(x_2, x_693); -x_696 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_697 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams), 10, 3); -lean_closure_set(x_697, 0, x_1); -lean_closure_set(x_697, 1, x_695); -lean_closure_set(x_697, 2, x_696); -x_698 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(x_427, x_697, x_4, x_5, x_6, x_7, x_8, x_9, x_694); -return x_698; +x_256 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingDomain___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__3(x_255, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +if (lean_obj_tag(x_256) == 0) +{ +lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; +x_257 = lean_ctor_get(x_256, 0); +lean_inc(x_257); +x_258 = lean_ctor_get(x_256, 1); +lean_inc(x_258); +lean_dec(x_256); +x_259 = lean_array_push(x_3, x_257); +x_260 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams), 10, 3); +lean_closure_set(x_260, 0, x_20); +lean_closure_set(x_260, 1, x_2); +lean_closure_set(x_260, 2, x_259); +x_261 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(x_15, x_260, x_5, x_6, x_7, x_8, x_9, x_10, x_258); +return x_261; } else { -uint8_t x_699; -lean_dec(x_427); +uint8_t x_262; +lean_dec(x_20); +lean_dec(x_15); +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_699 = !lean_is_exclusive(x_692); -if (x_699 == 0) +x_262 = !lean_is_exclusive(x_256); +if (x_262 == 0) { -return x_692; +return x_256; } else { -lean_object* x_700; lean_object* x_701; lean_object* x_702; -x_700 = lean_ctor_get(x_692, 0); -x_701 = lean_ctor_get(x_692, 1); -lean_inc(x_701); -lean_inc(x_700); -lean_dec(x_692); -x_702 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_702, 0, x_700); -lean_ctor_set(x_702, 1, x_701); -return x_702; +lean_object* x_263; lean_object* x_264; lean_object* x_265; +x_263 = lean_ctor_get(x_256, 0); +x_264 = lean_ctor_get(x_256, 1); +lean_inc(x_264); +lean_inc(x_263); +lean_dec(x_256); +x_265 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_265, 0, x_263); +lean_ctor_set(x_265, 1, x_264); +return x_265; } } } @@ -44238,593 +43229,32 @@ return x_702; } else { -lean_object* x_703; lean_object* x_704; lean_object* x_705; -lean_dec(x_428); -x_703 = lean_ctor_get(x_436, 1); -lean_inc(x_703); -lean_dec(x_436); -x_704 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams), 10, 3); -lean_closure_set(x_704, 0, x_1); -lean_closure_set(x_704, 1, x_2); -lean_closure_set(x_704, 2, x_435); -x_705 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(x_427, x_704, x_4, x_5, x_6, x_7, x_8, x_9, x_703); -return x_705; +lean_object* x_266; lean_object* x_267; +lean_dec(x_21); +lean_dec(x_16); +x_266 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux), 11, 4); +lean_closure_set(x_266, 0, x_20); +lean_closure_set(x_266, 1, x_2); +lean_closure_set(x_266, 2, x_3); +lean_closure_set(x_266, 3, x_23); +x_267 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(x_15, x_266, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +return x_267; } } else { -uint8_t x_706; -lean_dec(x_435); -lean_dec(x_428); -lean_dec(x_427); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_268; lean_object* x_269; lean_object* x_270; +lean_dec(x_13); lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_706 = !lean_is_exclusive(x_436); -if (x_706 == 0) -{ -return x_436; -} -else -{ -lean_object* x_707; lean_object* x_708; lean_object* x_709; -x_707 = lean_ctor_get(x_436, 0); -x_708 = lean_ctor_get(x_436, 1); -lean_inc(x_708); -lean_inc(x_707); -lean_dec(x_436); -x_709 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_709, 0, x_707); -lean_ctor_set(x_709, 1, x_708); -return x_709; -} -} -} -case 8: -{ -lean_object* x_710; lean_object* x_711; -lean_dec(x_12); -lean_dec(x_3); -x_710 = lean_ctor_get(x_11, 1); -lean_inc(x_710); -lean_dec(x_11); -lean_inc(x_8); -x_711 = l_Lean_PrettyPrinter_Delaborator_delab(x_4, x_5, x_6, x_7, x_8, x_9, x_710); -if (lean_obj_tag(x_711) == 0) -{ -uint8_t x_712; -x_712 = !lean_is_exclusive(x_711); -if (x_712 == 0) -{ -lean_object* x_713; lean_object* x_714; uint8_t x_715; lean_object* x_716; lean_object* x_717; size_t x_718; size_t x_719; lean_object* x_720; lean_object* x_721; lean_object* x_722; size_t x_723; lean_object* x_724; lean_object* x_725; lean_object* x_726; lean_object* x_727; lean_object* x_728; lean_object* x_729; lean_object* x_730; lean_object* x_731; lean_object* x_732; lean_object* x_733; lean_object* x_734; lean_object* x_735; lean_object* x_736; lean_object* x_737; -x_713 = lean_ctor_get(x_711, 0); -x_714 = lean_ctor_get(x_8, 5); -lean_inc(x_714); -lean_dec(x_8); -x_715 = 0; -x_716 = l_Lean_SourceInfo_fromRef(x_714, x_715); -x_717 = lean_array_get_size(x_2); -x_718 = lean_usize_of_nat(x_717); -lean_dec(x_717); -x_719 = 0; -x_720 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; -x_721 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__35(x_720, x_718, x_719, x_2); -x_722 = lean_array_get_size(x_721); -x_723 = lean_usize_of_nat(x_722); -lean_dec(x_722); -x_724 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__10; -x_725 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__36(x_724, x_723, x_719, x_721); -x_726 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_727 = l_Array_append___rarg(x_726, x_725); -x_728 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; -lean_inc(x_716); -x_729 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_729, 0, x_716); -lean_ctor_set(x_729, 1, x_728); -lean_ctor_set(x_729, 2, x_727); -x_730 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__8; -lean_inc(x_716); -x_731 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_731, 0, x_716); -lean_ctor_set(x_731, 1, x_730); -x_732 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__9; -lean_inc(x_716); -x_733 = l_Lean_Syntax_node2(x_716, x_732, x_731, x_713); -x_734 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__3; -lean_inc(x_716); -x_735 = l_Lean_Syntax_node2(x_716, x_734, x_729, x_733); -x_736 = l_Lean_PrettyPrinter_Delaborator_declSigWithId_formatter___closed__2; -x_737 = l_Lean_Syntax_node2(x_716, x_736, x_1, x_735); -lean_ctor_set(x_711, 0, x_737); -return x_711; -} -else -{ -lean_object* x_738; lean_object* x_739; lean_object* x_740; uint8_t x_741; lean_object* x_742; lean_object* x_743; size_t x_744; size_t x_745; lean_object* x_746; lean_object* x_747; lean_object* x_748; size_t x_749; lean_object* x_750; lean_object* x_751; lean_object* x_752; lean_object* x_753; lean_object* x_754; lean_object* x_755; lean_object* x_756; lean_object* x_757; lean_object* x_758; lean_object* x_759; lean_object* x_760; lean_object* x_761; lean_object* x_762; lean_object* x_763; lean_object* x_764; -x_738 = lean_ctor_get(x_711, 0); -x_739 = lean_ctor_get(x_711, 1); -lean_inc(x_739); -lean_inc(x_738); -lean_dec(x_711); -x_740 = lean_ctor_get(x_8, 5); -lean_inc(x_740); -lean_dec(x_8); -x_741 = 0; -x_742 = l_Lean_SourceInfo_fromRef(x_740, x_741); -x_743 = lean_array_get_size(x_2); -x_744 = lean_usize_of_nat(x_743); -lean_dec(x_743); -x_745 = 0; -x_746 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; -x_747 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__35(x_746, x_744, x_745, x_2); -x_748 = lean_array_get_size(x_747); -x_749 = lean_usize_of_nat(x_748); -lean_dec(x_748); -x_750 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__10; -x_751 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__36(x_750, x_749, x_745, x_747); -x_752 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_753 = l_Array_append___rarg(x_752, x_751); -x_754 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; -lean_inc(x_742); -x_755 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_755, 0, x_742); -lean_ctor_set(x_755, 1, x_754); -lean_ctor_set(x_755, 2, x_753); -x_756 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__8; -lean_inc(x_742); -x_757 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_757, 0, x_742); -lean_ctor_set(x_757, 1, x_756); -x_758 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__9; -lean_inc(x_742); -x_759 = l_Lean_Syntax_node2(x_742, x_758, x_757, x_738); -x_760 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__3; -lean_inc(x_742); -x_761 = l_Lean_Syntax_node2(x_742, x_760, x_755, x_759); -x_762 = l_Lean_PrettyPrinter_Delaborator_declSigWithId_formatter___closed__2; -x_763 = l_Lean_Syntax_node2(x_742, x_762, x_1, x_761); -x_764 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_764, 0, x_763); -lean_ctor_set(x_764, 1, x_739); -return x_764; -} -} -else -{ -uint8_t x_765; -lean_dec(x_8); -lean_dec(x_2); -lean_dec(x_1); -x_765 = !lean_is_exclusive(x_711); -if (x_765 == 0) -{ -return x_711; -} -else -{ -lean_object* x_766; lean_object* x_767; lean_object* x_768; -x_766 = lean_ctor_get(x_711, 0); -x_767 = lean_ctor_get(x_711, 1); -lean_inc(x_767); -lean_inc(x_766); -lean_dec(x_711); -x_768 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_768, 0, x_766); -lean_ctor_set(x_768, 1, x_767); -return x_768; -} -} -} -case 9: -{ -lean_object* x_769; lean_object* x_770; -lean_dec(x_12); -lean_dec(x_3); -x_769 = lean_ctor_get(x_11, 1); -lean_inc(x_769); -lean_dec(x_11); -lean_inc(x_8); -x_770 = l_Lean_PrettyPrinter_Delaborator_delab(x_4, x_5, x_6, x_7, x_8, x_9, x_769); -if (lean_obj_tag(x_770) == 0) -{ -uint8_t x_771; -x_771 = !lean_is_exclusive(x_770); -if (x_771 == 0) -{ -lean_object* x_772; lean_object* x_773; uint8_t x_774; lean_object* x_775; lean_object* x_776; size_t x_777; size_t x_778; lean_object* x_779; lean_object* x_780; lean_object* x_781; size_t x_782; lean_object* x_783; lean_object* x_784; lean_object* x_785; lean_object* x_786; lean_object* x_787; lean_object* x_788; lean_object* x_789; lean_object* x_790; lean_object* x_791; lean_object* x_792; lean_object* x_793; lean_object* x_794; lean_object* x_795; lean_object* x_796; -x_772 = lean_ctor_get(x_770, 0); -x_773 = lean_ctor_get(x_8, 5); -lean_inc(x_773); -lean_dec(x_8); -x_774 = 0; -x_775 = l_Lean_SourceInfo_fromRef(x_773, x_774); -x_776 = lean_array_get_size(x_2); -x_777 = lean_usize_of_nat(x_776); -lean_dec(x_776); -x_778 = 0; -x_779 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; -x_780 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__37(x_779, x_777, x_778, x_2); -x_781 = lean_array_get_size(x_780); -x_782 = lean_usize_of_nat(x_781); -lean_dec(x_781); -x_783 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__10; -x_784 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__38(x_783, x_782, x_778, x_780); -x_785 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_786 = l_Array_append___rarg(x_785, x_784); -x_787 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; -lean_inc(x_775); -x_788 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_788, 0, x_775); -lean_ctor_set(x_788, 1, x_787); -lean_ctor_set(x_788, 2, x_786); -x_789 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__8; -lean_inc(x_775); -x_790 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_790, 0, x_775); -lean_ctor_set(x_790, 1, x_789); -x_791 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__9; -lean_inc(x_775); -x_792 = l_Lean_Syntax_node2(x_775, x_791, x_790, x_772); -x_793 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__3; -lean_inc(x_775); -x_794 = l_Lean_Syntax_node2(x_775, x_793, x_788, x_792); -x_795 = l_Lean_PrettyPrinter_Delaborator_declSigWithId_formatter___closed__2; -x_796 = l_Lean_Syntax_node2(x_775, x_795, x_1, x_794); -lean_ctor_set(x_770, 0, x_796); -return x_770; -} -else -{ -lean_object* x_797; lean_object* x_798; lean_object* x_799; uint8_t x_800; lean_object* x_801; lean_object* x_802; size_t x_803; size_t x_804; lean_object* x_805; lean_object* x_806; lean_object* x_807; size_t x_808; lean_object* x_809; lean_object* x_810; lean_object* x_811; lean_object* x_812; lean_object* x_813; lean_object* x_814; lean_object* x_815; lean_object* x_816; lean_object* x_817; lean_object* x_818; lean_object* x_819; lean_object* x_820; lean_object* x_821; lean_object* x_822; lean_object* x_823; -x_797 = lean_ctor_get(x_770, 0); -x_798 = lean_ctor_get(x_770, 1); -lean_inc(x_798); -lean_inc(x_797); -lean_dec(x_770); -x_799 = lean_ctor_get(x_8, 5); -lean_inc(x_799); -lean_dec(x_8); -x_800 = 0; -x_801 = l_Lean_SourceInfo_fromRef(x_799, x_800); -x_802 = lean_array_get_size(x_2); -x_803 = lean_usize_of_nat(x_802); -lean_dec(x_802); -x_804 = 0; -x_805 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; -x_806 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__37(x_805, x_803, x_804, x_2); -x_807 = lean_array_get_size(x_806); -x_808 = lean_usize_of_nat(x_807); -lean_dec(x_807); -x_809 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__10; -x_810 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__38(x_809, x_808, x_804, x_806); -x_811 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_812 = l_Array_append___rarg(x_811, x_810); -x_813 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; -lean_inc(x_801); -x_814 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_814, 0, x_801); -lean_ctor_set(x_814, 1, x_813); -lean_ctor_set(x_814, 2, x_812); -x_815 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__8; -lean_inc(x_801); -x_816 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_816, 0, x_801); -lean_ctor_set(x_816, 1, x_815); -x_817 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__9; -lean_inc(x_801); -x_818 = l_Lean_Syntax_node2(x_801, x_817, x_816, x_797); -x_819 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__3; -lean_inc(x_801); -x_820 = l_Lean_Syntax_node2(x_801, x_819, x_814, x_818); -x_821 = l_Lean_PrettyPrinter_Delaborator_declSigWithId_formatter___closed__2; -x_822 = l_Lean_Syntax_node2(x_801, x_821, x_1, x_820); -x_823 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_823, 0, x_822); -lean_ctor_set(x_823, 1, x_798); -return x_823; -} -} -else -{ -uint8_t x_824; -lean_dec(x_8); -lean_dec(x_2); -lean_dec(x_1); -x_824 = !lean_is_exclusive(x_770); -if (x_824 == 0) -{ -return x_770; -} -else -{ -lean_object* x_825; lean_object* x_826; lean_object* x_827; -x_825 = lean_ctor_get(x_770, 0); -x_826 = lean_ctor_get(x_770, 1); -lean_inc(x_826); -lean_inc(x_825); -lean_dec(x_770); -x_827 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_827, 0, x_825); -lean_ctor_set(x_827, 1, x_826); -return x_827; -} -} -} -case 10: -{ -lean_object* x_828; lean_object* x_829; -lean_dec(x_12); lean_dec(x_3); -x_828 = lean_ctor_get(x_11, 1); -lean_inc(x_828); -lean_dec(x_11); -lean_inc(x_8); -x_829 = l_Lean_PrettyPrinter_Delaborator_delab(x_4, x_5, x_6, x_7, x_8, x_9, x_828); -if (lean_obj_tag(x_829) == 0) -{ -uint8_t x_830; -x_830 = !lean_is_exclusive(x_829); -if (x_830 == 0) -{ -lean_object* x_831; lean_object* x_832; uint8_t x_833; lean_object* x_834; lean_object* x_835; size_t x_836; size_t x_837; lean_object* x_838; lean_object* x_839; lean_object* x_840; size_t x_841; lean_object* x_842; lean_object* x_843; lean_object* x_844; lean_object* x_845; lean_object* x_846; lean_object* x_847; lean_object* x_848; lean_object* x_849; lean_object* x_850; lean_object* x_851; lean_object* x_852; lean_object* x_853; lean_object* x_854; lean_object* x_855; -x_831 = lean_ctor_get(x_829, 0); -x_832 = lean_ctor_get(x_8, 5); -lean_inc(x_832); -lean_dec(x_8); -x_833 = 0; -x_834 = l_Lean_SourceInfo_fromRef(x_832, x_833); -x_835 = lean_array_get_size(x_2); -x_836 = lean_usize_of_nat(x_835); -lean_dec(x_835); -x_837 = 0; -x_838 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; -x_839 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__39(x_838, x_836, x_837, x_2); -x_840 = lean_array_get_size(x_839); -x_841 = lean_usize_of_nat(x_840); -lean_dec(x_840); -x_842 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__10; -x_843 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__40(x_842, x_841, x_837, x_839); -x_844 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_845 = l_Array_append___rarg(x_844, x_843); -x_846 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; -lean_inc(x_834); -x_847 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_847, 0, x_834); -lean_ctor_set(x_847, 1, x_846); -lean_ctor_set(x_847, 2, x_845); -x_848 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__8; -lean_inc(x_834); -x_849 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_849, 0, x_834); -lean_ctor_set(x_849, 1, x_848); -x_850 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__9; -lean_inc(x_834); -x_851 = l_Lean_Syntax_node2(x_834, x_850, x_849, x_831); -x_852 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__3; -lean_inc(x_834); -x_853 = l_Lean_Syntax_node2(x_834, x_852, x_847, x_851); -x_854 = l_Lean_PrettyPrinter_Delaborator_declSigWithId_formatter___closed__2; -x_855 = l_Lean_Syntax_node2(x_834, x_854, x_1, x_853); -lean_ctor_set(x_829, 0, x_855); -return x_829; -} -else -{ -lean_object* x_856; lean_object* x_857; lean_object* x_858; uint8_t x_859; lean_object* x_860; lean_object* x_861; size_t x_862; size_t x_863; lean_object* x_864; lean_object* x_865; lean_object* x_866; size_t x_867; lean_object* x_868; lean_object* x_869; lean_object* x_870; lean_object* x_871; lean_object* x_872; lean_object* x_873; lean_object* x_874; lean_object* x_875; lean_object* x_876; lean_object* x_877; lean_object* x_878; lean_object* x_879; lean_object* x_880; lean_object* x_881; lean_object* x_882; -x_856 = lean_ctor_get(x_829, 0); -x_857 = lean_ctor_get(x_829, 1); -lean_inc(x_857); -lean_inc(x_856); -lean_dec(x_829); -x_858 = lean_ctor_get(x_8, 5); -lean_inc(x_858); -lean_dec(x_8); -x_859 = 0; -x_860 = l_Lean_SourceInfo_fromRef(x_858, x_859); -x_861 = lean_array_get_size(x_2); -x_862 = lean_usize_of_nat(x_861); -lean_dec(x_861); -x_863 = 0; -x_864 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; -x_865 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__39(x_864, x_862, x_863, x_2); -x_866 = lean_array_get_size(x_865); -x_867 = lean_usize_of_nat(x_866); -lean_dec(x_866); -x_868 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__10; -x_869 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__40(x_868, x_867, x_863, x_865); -x_870 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_871 = l_Array_append___rarg(x_870, x_869); -x_872 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; -lean_inc(x_860); -x_873 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_873, 0, x_860); -lean_ctor_set(x_873, 1, x_872); -lean_ctor_set(x_873, 2, x_871); -x_874 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__8; -lean_inc(x_860); -x_875 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_875, 0, x_860); -lean_ctor_set(x_875, 1, x_874); -x_876 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__9; -lean_inc(x_860); -x_877 = l_Lean_Syntax_node2(x_860, x_876, x_875, x_856); -x_878 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__3; -lean_inc(x_860); -x_879 = l_Lean_Syntax_node2(x_860, x_878, x_873, x_877); -x_880 = l_Lean_PrettyPrinter_Delaborator_declSigWithId_formatter___closed__2; -x_881 = l_Lean_Syntax_node2(x_860, x_880, x_1, x_879); -x_882 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_882, 0, x_881); -lean_ctor_set(x_882, 1, x_857); -return x_882; -} -} -else -{ -uint8_t x_883; -lean_dec(x_8); lean_dec(x_2); lean_dec(x_1); -x_883 = !lean_is_exclusive(x_829); -if (x_883 == 0) -{ -return x_829; -} -else -{ -lean_object* x_884; lean_object* x_885; lean_object* x_886; -x_884 = lean_ctor_get(x_829, 0); -x_885 = lean_ctor_get(x_829, 1); -lean_inc(x_885); -lean_inc(x_884); -lean_dec(x_829); -x_886 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_886, 0, x_884); -lean_ctor_set(x_886, 1, x_885); -return x_886; -} -} -} -default: -{ -lean_object* x_887; lean_object* x_888; +x_268 = lean_ctor_get(x_12, 1); +lean_inc(x_268); lean_dec(x_12); -lean_dec(x_3); -x_887 = lean_ctor_get(x_11, 1); -lean_inc(x_887); -lean_dec(x_11); -lean_inc(x_8); -x_888 = l_Lean_PrettyPrinter_Delaborator_delab(x_4, x_5, x_6, x_7, x_8, x_9, x_887); -if (lean_obj_tag(x_888) == 0) -{ -uint8_t x_889; -x_889 = !lean_is_exclusive(x_888); -if (x_889 == 0) -{ -lean_object* x_890; lean_object* x_891; uint8_t x_892; lean_object* x_893; lean_object* x_894; size_t x_895; size_t x_896; lean_object* x_897; lean_object* x_898; lean_object* x_899; size_t x_900; lean_object* x_901; lean_object* x_902; lean_object* x_903; lean_object* x_904; lean_object* x_905; lean_object* x_906; lean_object* x_907; lean_object* x_908; lean_object* x_909; lean_object* x_910; lean_object* x_911; lean_object* x_912; lean_object* x_913; lean_object* x_914; -x_890 = lean_ctor_get(x_888, 0); -x_891 = lean_ctor_get(x_8, 5); -lean_inc(x_891); -lean_dec(x_8); -x_892 = 0; -x_893 = l_Lean_SourceInfo_fromRef(x_891, x_892); -x_894 = lean_array_get_size(x_2); -x_895 = lean_usize_of_nat(x_894); -lean_dec(x_894); -x_896 = 0; -x_897 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; -x_898 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__41(x_897, x_895, x_896, x_2); -x_899 = lean_array_get_size(x_898); -x_900 = lean_usize_of_nat(x_899); -lean_dec(x_899); -x_901 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__10; -x_902 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__42(x_901, x_900, x_896, x_898); -x_903 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_904 = l_Array_append___rarg(x_903, x_902); -x_905 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; -lean_inc(x_893); -x_906 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_906, 0, x_893); -lean_ctor_set(x_906, 1, x_905); -lean_ctor_set(x_906, 2, x_904); -x_907 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__8; -lean_inc(x_893); -x_908 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_908, 0, x_893); -lean_ctor_set(x_908, 1, x_907); -x_909 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__9; -lean_inc(x_893); -x_910 = l_Lean_Syntax_node2(x_893, x_909, x_908, x_890); -x_911 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__3; -lean_inc(x_893); -x_912 = l_Lean_Syntax_node2(x_893, x_911, x_906, x_910); -x_913 = l_Lean_PrettyPrinter_Delaborator_declSigWithId_formatter___closed__2; -x_914 = l_Lean_Syntax_node2(x_893, x_913, x_1, x_912); -lean_ctor_set(x_888, 0, x_914); -return x_888; -} -else -{ -lean_object* x_915; lean_object* x_916; lean_object* x_917; uint8_t x_918; lean_object* x_919; lean_object* x_920; size_t x_921; size_t x_922; lean_object* x_923; lean_object* x_924; lean_object* x_925; size_t x_926; lean_object* x_927; lean_object* x_928; lean_object* x_929; lean_object* x_930; lean_object* x_931; lean_object* x_932; lean_object* x_933; lean_object* x_934; lean_object* x_935; lean_object* x_936; lean_object* x_937; lean_object* x_938; lean_object* x_939; lean_object* x_940; lean_object* x_941; -x_915 = lean_ctor_get(x_888, 0); -x_916 = lean_ctor_get(x_888, 1); -lean_inc(x_916); -lean_inc(x_915); -lean_dec(x_888); -x_917 = lean_ctor_get(x_8, 5); -lean_inc(x_917); -lean_dec(x_8); -x_918 = 0; -x_919 = l_Lean_SourceInfo_fromRef(x_917, x_918); -x_920 = lean_array_get_size(x_2); -x_921 = lean_usize_of_nat(x_920); -lean_dec(x_920); -x_922 = 0; -x_923 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; -x_924 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__41(x_923, x_921, x_922, x_2); -x_925 = lean_array_get_size(x_924); -x_926 = lean_usize_of_nat(x_925); -lean_dec(x_925); -x_927 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__10; -x_928 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__42(x_927, x_926, x_922, x_924); -x_929 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_930 = l_Array_append___rarg(x_929, x_928); -x_931 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; -lean_inc(x_919); -x_932 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_932, 0, x_919); -lean_ctor_set(x_932, 1, x_931); -lean_ctor_set(x_932, 2, x_930); -x_933 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__8; -lean_inc(x_919); -x_934 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_934, 0, x_919); -lean_ctor_set(x_934, 1, x_933); -x_935 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__9; -lean_inc(x_919); -x_936 = l_Lean_Syntax_node2(x_919, x_935, x_934, x_915); -x_937 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__3; -lean_inc(x_919); -x_938 = l_Lean_Syntax_node2(x_919, x_937, x_932, x_936); -x_939 = l_Lean_PrettyPrinter_Delaborator_declSigWithId_formatter___closed__2; -x_940 = l_Lean_Syntax_node2(x_919, x_939, x_1, x_938); -x_941 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_941, 0, x_940); -lean_ctor_set(x_941, 1, x_916); -return x_941; -} -} -else -{ -uint8_t x_942; -lean_dec(x_8); -lean_dec(x_2); -lean_dec(x_1); -x_942 = !lean_is_exclusive(x_888); -if (x_942 == 0) -{ -return x_888; -} -else -{ -lean_object* x_943; lean_object* x_944; lean_object* x_945; -x_943 = lean_ctor_get(x_888, 0); -x_944 = lean_ctor_get(x_888, 1); -lean_inc(x_944); -lean_inc(x_943); -lean_dec(x_888); -x_945 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_945, 0, x_943); -lean_ctor_set(x_945, 1, x_944); -return x_945; -} -} -} +x_269 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___closed__2; +x_270 = l_panic___at_Lean_PrettyPrinter_Delaborator_delabFVar___spec__1(x_269, x_5, x_6, x_7, x_8, x_9, x_10, x_268); +return x_270; } } } @@ -44854,253 +43284,12 @@ lean_dec(x_1); return x_7; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__4(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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__5(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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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; -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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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; -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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__7(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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__8___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__8(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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__9___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__9(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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__10___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__10(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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__11___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__11(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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__12___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__12(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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__13___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__13(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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__15___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_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__15(x_1, x_2, x_6, x_7, x_5); -lean_dec(x_2); -lean_dec(x_1); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__18___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_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__18(x_1, x_2, x_6, x_7, x_5); -lean_dec(x_2); -lean_dec(x_1); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__19___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_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__19(x_1, x_2, x_6, x_7, x_5); -lean_dec(x_2); -lean_dec(x_1); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__20___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_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__20(x_1, x_2, x_6, x_7, x_5); -lean_dec(x_2); -lean_dec(x_1); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__21___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_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__21(x_1, x_2, x_6, x_7, x_5); -lean_dec(x_2); -lean_dec(x_1); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__23___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -x_9 = l_Lean_throwError___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__23(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Lean_ofExcept___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_ofExcept___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__22(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_9 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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); @@ -45108,7 +43297,7 @@ lean_dec(x_2); return x_9; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__24___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__1___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; @@ -45116,13 +43305,13 @@ x_6 = lean_unbox_usize(x_3); lean_dec(x_3); x_7 = lean_unbox_usize(x_4); lean_dec(x_4); -x_8 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__24(x_1, x_2, x_6, x_7, x_5); +x_8 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__1(x_1, x_2, x_6, x_7, x_5); lean_dec(x_2); lean_dec(x_1); return x_8; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__25___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__2___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; @@ -45130,13 +43319,13 @@ x_6 = lean_unbox_usize(x_3); lean_dec(x_3); x_7 = lean_unbox_usize(x_4); lean_dec(x_4); -x_8 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__25(x_1, x_2, x_6, x_7, x_5); +x_8 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__2(x_1, x_2, x_6, x_7, x_5); lean_dec(x_2); lean_dec(x_1); return x_8; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__26___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__3___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; @@ -45144,13 +43333,13 @@ x_6 = lean_unbox_usize(x_3); lean_dec(x_3); x_7 = lean_unbox_usize(x_4); lean_dec(x_4); -x_8 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__26(x_1, x_2, x_6, x_7, x_5); +x_8 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__3(x_1, x_2, x_6, x_7, x_5); lean_dec(x_2); lean_dec(x_1); return x_8; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__27___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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; @@ -45158,13 +43347,13 @@ x_6 = lean_unbox_usize(x_3); lean_dec(x_3); x_7 = lean_unbox_usize(x_4); lean_dec(x_4); -x_8 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__27(x_1, x_2, x_6, x_7, x_5); +x_8 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__4(x_1, x_2, x_6, x_7, x_5); lean_dec(x_2); lean_dec(x_1); return x_8; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__28___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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; @@ -45172,27 +43361,41 @@ x_6 = lean_unbox_usize(x_3); lean_dec(x_3); x_7 = lean_unbox_usize(x_4); lean_dec(x_4); -x_8 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__28(x_1, x_2, x_6, x_7, x_5); +x_8 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__5(x_1, x_2, x_6, x_7, x_5); lean_dec(x_2); lean_dec(x_1); return x_8; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__29___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_throwError___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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) { _start: { -size_t x_6; size_t x_7; lean_object* x_8; -x_6 = lean_unbox_usize(x_3); +lean_object* x_9; +x_9 = l_Lean_throwError___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__7(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_7 = lean_unbox_usize(x_4); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_ofExcept___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_ofExcept___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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); lean_dec(x_4); -x_8 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__29(x_1, x_2, x_6, x_7, x_5); +lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -return x_8; +return x_9; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__30___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__8___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; @@ -45200,13 +43403,13 @@ x_6 = lean_unbox_usize(x_3); lean_dec(x_3); x_7 = lean_unbox_usize(x_4); lean_dec(x_4); -x_8 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__30(x_1, x_2, x_6, x_7, x_5); +x_8 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__8(x_1, x_2, x_6, x_7, x_5); lean_dec(x_2); lean_dec(x_1); return x_8; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__31___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__9___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; @@ -45214,13 +43417,13 @@ x_6 = lean_unbox_usize(x_3); lean_dec(x_3); x_7 = lean_unbox_usize(x_4); lean_dec(x_4); -x_8 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__31(x_1, x_2, x_6, x_7, x_5); +x_8 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__9(x_1, x_2, x_6, x_7, x_5); lean_dec(x_2); lean_dec(x_1); return x_8; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__32___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__10___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; @@ -45228,13 +43431,13 @@ x_6 = lean_unbox_usize(x_3); lean_dec(x_3); x_7 = lean_unbox_usize(x_4); lean_dec(x_4); -x_8 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__32(x_1, x_2, x_6, x_7, x_5); +x_8 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__10(x_1, x_2, x_6, x_7, x_5); lean_dec(x_2); lean_dec(x_1); return x_8; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__33___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__11___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; @@ -45242,13 +43445,13 @@ x_6 = lean_unbox_usize(x_3); lean_dec(x_3); x_7 = lean_unbox_usize(x_4); lean_dec(x_4); -x_8 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__33(x_1, x_2, x_6, x_7, x_5); +x_8 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__11(x_1, x_2, x_6, x_7, x_5); lean_dec(x_2); lean_dec(x_1); return x_8; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__34___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { size_t x_6; size_t x_7; lean_object* x_8; @@ -45256,121 +43459,101 @@ x_6 = lean_unbox_usize(x_3); lean_dec(x_3); x_7 = lean_unbox_usize(x_4); lean_dec(x_4); -x_8 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__34(x_1, x_2, x_6, x_7, x_5); +x_8 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__12(x_1, x_2, x_6, x_7, x_5); lean_dec(x_2); lean_dec(x_1); return x_8; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__35___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__13___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_5; size_t x_6; lean_object* x_7; -x_5 = lean_unbox_usize(x_2); -lean_dec(x_2); +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 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__35(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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__36___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); +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__13(x_1, x_2, x_6, x_7, x_5); lean_dec(x_2); -x_6 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_7 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__36(x_1, x_5, x_6, x_4); lean_dec(x_1); -return x_7; +return x_8; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__37___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__14___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_5; size_t x_6; lean_object* x_7; -x_5 = lean_unbox_usize(x_2); -lean_dec(x_2); +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 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__37(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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__38___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); +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__14(x_1, x_2, x_6, x_7, x_5); lean_dec(x_2); -x_6 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_7 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__38(x_1, x_5, x_6, x_4); lean_dec(x_1); -return x_7; +return x_8; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__39___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__15___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_5; size_t x_6; lean_object* x_7; -x_5 = lean_unbox_usize(x_2); -lean_dec(x_2); +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 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__39(x_1, x_5, x_6, x_4); +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__15(x_1, x_2, x_6, x_7, x_5); +lean_dec(x_2); lean_dec(x_1); -return x_7; +return x_8; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__40___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__16___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_5; size_t x_6; lean_object* x_7; -x_5 = lean_unbox_usize(x_2); -lean_dec(x_2); +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 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__40(x_1, x_5, x_6, x_4); +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__16(x_1, x_2, x_6, x_7, x_5); +lean_dec(x_2); lean_dec(x_1); -return x_7; +return x_8; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__17___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_5; size_t x_6; lean_object* x_7; -x_5 = lean_unbox_usize(x_2); -lean_dec(x_2); +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 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__41(x_1, x_5, x_6, x_4); +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__17(x_1, x_2, x_6, x_7, x_5); +lean_dec(x_2); lean_dec(x_1); -return x_7; +return x_8; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__42___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__18___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_5; size_t x_6; lean_object* x_7; -x_5 = lean_unbox_usize(x_2); -lean_dec(x_2); +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 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__42(x_1, x_5, x_6, x_4); +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__18(x_1, x_2, x_6, x_7, x_5); +lean_dec(x_2); lean_dec(x_1); -return x_7; +return x_8; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_10); lean_dec(x_8); lean_dec(x_7); @@ -45380,11 +43563,11 @@ lean_dec(x_3); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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_8); lean_dec(x_7); @@ -45394,11 +43577,11 @@ lean_dec(x_3); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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) { _start: { lean_object* x_12; -x_12 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_10); lean_dec(x_8); lean_dec(x_7); @@ -45408,11 +43591,11 @@ lean_dec(x_3); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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_12; -x_12 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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_10); lean_dec(x_8); lean_dec(x_7); @@ -45422,11 +43605,11 @@ lean_dec(x_3); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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_10); lean_dec(x_8); lean_dec(x_7); @@ -45436,11 +43619,11 @@ lean_dec(x_3); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +x_8 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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); @@ -45450,20 +43633,20 @@ lean_dec(x_1); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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) { _start: { -lean_object* x_14; -x_14 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_12, x_13); -lean_dec(x_5); -return x_14; +lean_object* x_13; +x_13 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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); +lean_dec(x_4); +return x_13; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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) { _start: { lean_object* x_12; -x_12 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_10); lean_dec(x_8); lean_dec(x_7); @@ -45473,11 +43656,11 @@ lean_dec(x_3); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_12; -x_12 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__10(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_12 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_10); lean_dec(x_8); lean_dec(x_7); @@ -45487,11 +43670,11 @@ lean_dec(x_3); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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) { _start: { lean_object* x_12; -x_12 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__11(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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__10(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_10); lean_dec(x_8); lean_dec(x_7); @@ -45501,11 +43684,11 @@ lean_dec(x_3); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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) { _start: { lean_object* x_12; -x_12 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__12(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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__11(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_10); lean_dec(x_8); lean_dec(x_7); @@ -45515,11 +43698,11 @@ lean_dec(x_3); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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) { _start: { lean_object* x_12; -x_12 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___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 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__12(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_10); lean_dec(x_8); lean_dec(x_7); @@ -45529,11 +43712,11 @@ lean_dec(x_3); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__14___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___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) { _start: { lean_object* x_12; -x_12 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__14(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_12 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__13(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_10); lean_dec(x_8); lean_dec(x_7); @@ -45543,11 +43726,11 @@ lean_dec(x_3); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__14___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_12; -x_12 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__15(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_12 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__14(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_10); lean_dec(x_8); lean_dec(x_7); @@ -45557,20 +43740,20 @@ lean_dec(x_3); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__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_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__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) { _start: { -lean_object* x_13; -x_13 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__16(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_4); -return x_13; +lean_object* x_12; +x_12 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__15(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_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__17___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__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) { _start: { lean_object* x_12; -x_12 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__17(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_12 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__16(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_10); lean_dec(x_8); lean_dec(x_7); @@ -45580,11 +43763,11 @@ lean_dec(x_3); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__18___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__17___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_12; -x_12 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__18(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_12 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__17(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_10); lean_dec(x_8); lean_dec(x_7); @@ -45594,17 +43777,16 @@ lean_dec(x_3); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__18___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; -x_10 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__19(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_10 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__18(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_1); return x_10; } } @@ -45624,39 +43806,40 @@ x_1 = l_Lean_pp_fullNames; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___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_PrettyPrinter_Delaborator_delabConstWithSignature___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, lean_object* x_8) { _start: { -uint8_t x_8; -x_8 = !lean_is_exclusive(x_5); -if (x_8 == 0) +uint8_t x_9; +x_9 = !lean_is_exclusive(x_6); +if (x_9 == 0) { -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; -x_9 = lean_ctor_get(x_5, 2); -x_10 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__1___closed__1; -x_11 = 1; -x_12 = l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(x_9, x_10, x_11); +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; +x_10 = lean_ctor_get(x_6, 2); +x_11 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__1___closed__1; +x_12 = l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(x_10, x_11, x_1); x_13 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__1___closed__2; -x_14 = l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(x_12, x_13, x_11); -lean_ctor_set(x_5, 2, x_14); -x_15 = l_Lean_PrettyPrinter_Delaborator_delabConst(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -return x_15; +x_14 = 1; +x_15 = l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(x_12, x_13, x_14); +lean_ctor_set(x_6, 2, x_15); +x_16 = l_Lean_PrettyPrinter_Delaborator_delabConst(x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_16; } 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; uint8_t x_27; lean_object* x_28; uint8_t x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_16 = lean_ctor_get(x_5, 0); -x_17 = lean_ctor_get(x_5, 1); -x_18 = lean_ctor_get(x_5, 2); -x_19 = lean_ctor_get(x_5, 3); -x_20 = lean_ctor_get(x_5, 4); -x_21 = lean_ctor_get(x_5, 5); -x_22 = lean_ctor_get(x_5, 6); -x_23 = lean_ctor_get(x_5, 7); -x_24 = lean_ctor_get(x_5, 8); -x_25 = lean_ctor_get(x_5, 9); -x_26 = lean_ctor_get(x_5, 10); -x_27 = lean_ctor_get_uint8(x_5, sizeof(void*)*11); +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; uint8_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_17 = lean_ctor_get(x_6, 0); +x_18 = lean_ctor_get(x_6, 1); +x_19 = lean_ctor_get(x_6, 2); +x_20 = lean_ctor_get(x_6, 3); +x_21 = lean_ctor_get(x_6, 4); +x_22 = lean_ctor_get(x_6, 5); +x_23 = lean_ctor_get(x_6, 6); +x_24 = lean_ctor_get(x_6, 7); +x_25 = lean_ctor_get(x_6, 8); +x_26 = lean_ctor_get(x_6, 9); +x_27 = lean_ctor_get(x_6, 10); +x_28 = lean_ctor_get_uint8(x_6, sizeof(void*)*11); +lean_inc(x_27); lean_inc(x_26); lean_inc(x_25); lean_inc(x_24); @@ -45667,150 +43850,164 @@ lean_inc(x_20); lean_inc(x_19); lean_inc(x_18); lean_inc(x_17); -lean_inc(x_16); -lean_dec(x_5); -x_28 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__1___closed__1; -x_29 = 1; -x_30 = l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(x_18, x_28, x_29); +lean_dec(x_6); +x_29 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__1___closed__1; +x_30 = l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(x_19, x_29, x_1); x_31 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__1___closed__2; -x_32 = l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(x_30, x_31, x_29); -x_33 = lean_alloc_ctor(0, 11, 1); -lean_ctor_set(x_33, 0, x_16); -lean_ctor_set(x_33, 1, x_17); -lean_ctor_set(x_33, 2, x_32); -lean_ctor_set(x_33, 3, x_19); -lean_ctor_set(x_33, 4, x_20); -lean_ctor_set(x_33, 5, x_21); -lean_ctor_set(x_33, 6, x_22); -lean_ctor_set(x_33, 7, x_23); -lean_ctor_set(x_33, 8, x_24); -lean_ctor_set(x_33, 9, x_25); -lean_ctor_set(x_33, 10, x_26); -lean_ctor_set_uint8(x_33, sizeof(void*)*11, x_27); -x_34 = l_Lean_PrettyPrinter_Delaborator_delabConst(x_1, x_2, x_3, x_4, x_33, x_6, x_7); -return x_34; -} -} +x_32 = 1; +x_33 = l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(x_30, x_31, x_32); +x_34 = lean_alloc_ctor(0, 11, 1); +lean_ctor_set(x_34, 0, x_17); +lean_ctor_set(x_34, 1, x_18); +lean_ctor_set(x_34, 2, x_33); +lean_ctor_set(x_34, 3, x_20); +lean_ctor_set(x_34, 4, x_21); +lean_ctor_set(x_34, 5, x_22); +lean_ctor_set(x_34, 6, x_23); +lean_ctor_set(x_34, 7, x_24); +lean_ctor_set(x_34, 8, x_25); +lean_ctor_set(x_34, 9, x_26); +lean_ctor_set(x_34, 10, x_27); +lean_ctor_set_uint8(x_34, sizeof(void*)*11, x_28); +x_35 = l_Lean_PrettyPrinter_Delaborator_delabConst(x_2, x_3, x_4, x_5, x_34, x_7, x_8); +return x_35; } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__1), 7, 0); -return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature(lean_object* x_1, lean_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_PrettyPrinter_Delaborator_delabConstWithSignature(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: { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_8 = l_Lean_PrettyPrinter_Delaborator_SubExpr_getExpr___at_Lean_PrettyPrinter_Delaborator_getExprKind___spec__1(x_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); +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_9 = l_Lean_PrettyPrinter_Delaborator_SubExpr_getExpr___at_Lean_PrettyPrinter_Delaborator_getExprKind___spec__1(x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_10 = lean_ctor_get(x_9, 0); lean_inc(x_10); -lean_dec(x_8); -x_11 = lean_unsigned_to_nat(0u); -x_12 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___closed__1; +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +x_12 = lean_box(x_1); +x_13 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__1___boxed), 8, 1); +lean_closure_set(x_13, 0, x_12); +x_14 = lean_unsigned_to_nat(0u); +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_1); -lean_inc(x_9); -x_13 = l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delab___spec__3(x_9, x_11, x_12, x_1, x_2, x_3, x_4, x_5, x_6, x_10); -if (lean_obj_tag(x_13) == 0) +lean_inc(x_10); +x_15 = l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delab___spec__3(x_10, x_14, x_13, x_2, x_3, x_4, x_5, x_6, x_7, x_11); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_3); -x_16 = lean_infer_type(x_9, x_3, x_4, x_5, x_6, x_15); -if (lean_obj_tag(x_16) == 0) +x_18 = lean_infer_type(x_10, x_4, x_5, x_6, x_7, x_17); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); -lean_inc(x_18); -lean_dec(x_16); -x_19 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_20 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams), 10, 3); -lean_closure_set(x_20, 0, x_14); -lean_closure_set(x_20, 1, x_19); -lean_closure_set(x_20, 2, x_19); -x_21 = lean_unsigned_to_nat(1u); -x_22 = l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delab___spec__3(x_17, x_21, x_20, x_1, x_2, x_3, x_4, x_5, x_6, x_18); -return x_22; +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_19 = 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_NameSet_empty; +x_22 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; +x_23 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams), 10, 3); +lean_closure_set(x_23, 0, x_21); +lean_closure_set(x_23, 1, x_16); +lean_closure_set(x_23, 2, x_22); +x_24 = lean_unsigned_to_nat(1u); +x_25 = l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delab___spec__3(x_19, x_24, x_23, x_2, x_3, x_4, x_5, x_6, x_7, x_20); +return x_25; } else { -uint8_t x_23; -lean_dec(x_14); +uint8_t x_26; +lean_dec(x_16); +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_23 = !lean_is_exclusive(x_16); -if (x_23 == 0) +x_26 = !lean_is_exclusive(x_18); +if (x_26 == 0) { -return x_16; +return x_18; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_16, 0); -x_25 = lean_ctor_get(x_16, 1); -lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_16); -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_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_18, 0); +x_28 = lean_ctor_get(x_18, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_18); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; } } } else { -uint8_t x_27; -lean_dec(x_9); +uint8_t x_30; +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); -lean_dec(x_1); -x_27 = !lean_is_exclusive(x_13); -if (x_27 == 0) +x_30 = !lean_is_exclusive(x_15); +if (x_30 == 0) { -return x_13; +return x_15; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_13, 0); -x_29 = lean_ctor_get(x_13, 1); -lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_13); -x_30 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_30, 0, x_28); -lean_ctor_set(x_30, 1, x_29); -return x_30; +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_15, 0); +x_32 = lean_ctor_get(x_15, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_15); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; } } } } +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___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_9; lean_object* x_10; +x_9 = lean_unbox(x_1); +lean_dec(x_1); +x_10 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__1(x_9, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___boxed(lean_object* x_1, lean_object* x_2, 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_Lean_PrettyPrinter_Delaborator_delabConstWithSignature(x_9, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_10; +} +} lean_object* initialize_Lean_Parser(uint8_t builtin, lean_object*); lean_object* initialize_Lean_PrettyPrinter_Delaborator_Basic(uint8_t builtin, lean_object*); lean_object* initialize_Lean_PrettyPrinter_Delaborator_SubExpr(uint8_t builtin, lean_object*); @@ -47018,18 +45215,8 @@ l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delabo lean_mark_persistent(l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_evalSyntaxConstantUnsafe___closed__1); l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_evalSyntaxConstantUnsafe___closed__2 = _init_l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_evalSyntaxConstantUnsafe___closed__2(); lean_mark_persistent(l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_evalSyntaxConstantUnsafe___closed__2); -l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__1___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__1___closed__1(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__1___closed__1); -l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__8___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__8___closed__1(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__8___closed__1); -l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__8___closed__2 = _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__8___closed__2(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__8___closed__2); -l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__8___closed__3 = _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__8___closed__3(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__8___closed__3); -l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__16___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__16___closed__1(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__16___closed__1); -l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__16___closed__2 = _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__16___closed__2(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___lambda__16___closed__2); +l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__1(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__1); l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__1(); lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__1); l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__2 = _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__2(); @@ -47056,14 +45243,30 @@ l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__1 lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__12); l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__13 = _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__13(); lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__13); -l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__14 = _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__14(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__14); +l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__7___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__7___closed__1(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__7___closed__1); +l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__7___closed__2 = _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__7___closed__2(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__7___closed__2); +l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__7___closed__3 = _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__7___closed__3(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__7___closed__3); +l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__7___closed__4 = _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__7___closed__4(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__7___closed__4); +l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__15___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__15___closed__1(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__15___closed__1); +l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__15___closed__2 = _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__15___closed__2(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__15___closed__2); +l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___closed__1(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___closed__1); +l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___closed__2 = _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___closed__2(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___closed__2); +l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___closed__3 = _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___closed__3(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___closed__3); +l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___closed__4 = _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___closed__4(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___closed__4); l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__1___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__1___closed__1(); lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__1___closed__1); l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__1___closed__2 = _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__1___closed__2(); lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__1___closed__2); -l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___closed__1(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___closed__1); 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 27f3dd9adfa7..9e1a6e735b47 100644 --- a/stage0/stdlib/Lean/Util/Trace.c +++ b/stage0/stdlib/Lean/Util/Trace.c @@ -13,7 +13,6 @@ #ifdef __cplusplus extern "C" { #endif -LEAN_EXPORT 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_exceptOptionEmoji___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___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*); LEAN_EXPORT lean_object* l_Lean_instMonadAlwaysExceptReaderT___rarg(lean_object*); @@ -26,9 +25,9 @@ static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__22; LEAN_EXPORT lean_object* l_Lean_addTrace___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg___lambda__4___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_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; @@ -36,47 +35,45 @@ static lean_object* l_Array_forInUnsafe_loop___at_Lean_printTraces___spec__4___r 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*); -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2607____closed__18; 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*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, lean_object*); +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*, uint8_t, lean_object*, lean_object*, lean_object*, uint8_t, double, 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*); +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_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_printTraces___spec__1(lean_object*); 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_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2601____closed__21; LEAN_EXPORT lean_object* l_Lean_shouldEnableNestedTrace___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1421____closed__1; 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*, uint8_t, 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_Lean_withTraceNode___rarg___lambda__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__2; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___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_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*); 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*, uint8_t, 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*, uint8_t, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); double lean_float_div(double, double); +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2601____closed__24; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2607____closed__7; 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_2607____closed__4; 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*); +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2601____closed__3; static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__31; -static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__2; LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_withTraceNode___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1421____closed__2; LEAN_EXPORT lean_object* l_Lean_resetTraceState___rarg___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_modifyTraces___rarg(lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2607____closed__26; 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___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instMonadAlwaysExceptStateRefT_x27(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*, lean_object*, uint8_t, lean_object*); @@ -90,10 +87,10 @@ 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*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_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*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, lean_object*, lean_object*, lean_object*); static double l_Lean_withSeconds___rarg___lambda__1___closed__1; -static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__5; LEAN_EXPORT lean_object* l_Lean_doElemTrace_x5b___x5d____; 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; @@ -104,15 +101,15 @@ LEAN_EXPORT lean_object* l_Lean_printTraces___rarg(lean_object*, lean_object*, l static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__48; LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_withTraceNode_x27___spec__3(lean_object*, lean_object*); 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*); +LEAN_EXPORT lean_object* l___auto____x40_Lean_Util_Trace___hyg_2601_; lean_object* l_StateT_instMonadExceptOfStateT___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__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*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__9; LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__6___boxed(lean_object*, lean_object*); static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__5; -static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__4; 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_object*, lean_object*, lean_object*, 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__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_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; @@ -134,6 +131,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___ra LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_withTraceNode___spec__1(lean_object*, lean_object*); 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*); +static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1415____closed__3; 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; @@ -142,27 +140,28 @@ 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*); +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2601____closed__10; LEAN_EXPORT lean_object* l_Lean_instInhabitedTraceState; +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2601____closed__14; 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_MonadExcept_ofExcept___at_Lean_withTraceNodeBefore___spec__2(lean_object*, lean_object*); -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*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*); +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, lean_object*, 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_2607____closed__13; lean_object* l_instMonadExceptOfEIO(lean_object*); -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*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__3___boxed(lean_object**); static lean_object* l_Lean_instInhabitedTraceElem___closed__1; -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2607____closed__5; static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__36; LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_withTraceNode_x27___spec__3___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2601____closed__7; 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; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNode___spec__1___boxed(lean_object*, lean_object*, lean_object*); 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*); @@ -175,47 +174,52 @@ 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_2601____closed__16; 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; 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*, uint8_t, lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__5(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_object*, lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__3; LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___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_object*, 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*); static lean_object* l_Lean_withTraceNode___rarg___lambda__6___closed__2; +static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__4; +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2601____closed__26; lean_object* lean_st_ref_take(lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2607____closed__20; LEAN_EXPORT lean_object* l_Lean_instMonadTrace___rarg___lambda__1(lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2601____closed__6; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___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_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*, uint8_t, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_withTraceNodeBefore___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__3; static lean_object* l_Lean_resetTraceState___rarg___closed__1; LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_withTraceNode___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2607____closed__15; LEAN_EXPORT lean_object* l_Lean_setTraceState___rarg___lambda__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg___lambda__3(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_exceptEmoji___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___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*, lean_object*); +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2601____closed__15; 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*); lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); static lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__3___closed__1; +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2601____closed__23; 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*); 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*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2607____closed__16; +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2601____closed__1; +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*, 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___auto____x40_Lean_Util_Trace___hyg_2601____closed__19; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__3(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_object*); 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*); -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2607____closed__14; +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2601____closed__22; 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*); static lean_object* l_Lean_withTraceNode___rarg___lambda__6___closed__1; lean_object* l_Lean_Syntax_getKind(lean_object*); @@ -223,6 +227,7 @@ LEAN_EXPORT lean_object* l_Lean_resetTraceState___rarg___lambda__1___boxed(lean_ 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; static lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__3___closed__2; +static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__3; 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*); lean_object* l_ReaderT_instMonadExceptOfReaderT___rarg(lean_object*); @@ -235,7 +240,6 @@ LEAN_EXPORT lean_object* l_Lean_bombEmoji; LEAN_EXPORT lean_object* l_Lean_checkEmoji; lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___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_object*); -LEAN_EXPORT lean_object* l___auto____x40_Lean_Util_Trace___hyg_2607_; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__5___boxed(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*); @@ -255,6 +259,7 @@ LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__47; 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_EXPORT lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1415_(lean_object*); 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_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*); @@ -276,25 +281,20 @@ static lean_object* l_Lean_withTraceNode___rarg___lambda__5___closed__1; LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__14(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_object*, lean_object*); 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; -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2607____closed__24; -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2607____closed__6; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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_2607____closed__19; 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_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*); +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2601____closed__11; 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*); -static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____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_MonadExcept_ofExcept___at_Lean_withTraceNode___spec__2(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___auto____x40_Lean_Util_Trace___hyg_2607____closed__3; 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; LEAN_EXPORT lean_object* l_Lean_withTraceNode_x27___rarg___lambda__2(lean_object*); @@ -303,6 +303,7 @@ static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doEle 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*); +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2601____closed__18; 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); @@ -312,19 +313,21 @@ LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___rarg___lambda__3(lean_obje 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_initFn____x40_Lean_Util_Trace___hyg_1415____closed__2; 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_2607____closed__17; -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2607____closed__27; +static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__5; lean_object* lean_register_option(lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2601____closed__25; 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*); -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2607____closed__12; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2607____closed__1; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNode___spec__1(size_t, size_t, lean_object*); +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2601____closed__8; +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2601____closed__27; 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_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__6; 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; @@ -343,8 +346,9 @@ 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_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_Lean_exceptEmoji___rarg(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*); -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1330_(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*); +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2601____closed__17; +static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__1; LEAN_EXPORT lean_object* l_Lean_printTraces___rarg___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, uint8_t); static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__11; @@ -356,8 +360,8 @@ static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__10; lean_object* l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(lean_object*, 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___auto____x40_Lean_Util_Trace___hyg_2601____closed__2; 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_initFn____x40_Lean_Util_Trace___hyg_1421_(lean_object*); LEAN_EXPORT lean_object* l_Lean_instMonadAlwaysExceptEIO(lean_object*); 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*, lean_object*); LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_withTraceNode_x27___spec__2(lean_object*, lean_object*); @@ -366,7 +370,6 @@ static lean_object* l___private_Lean_Util_Trace_0__Lean_checkTraceOption___close 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*); lean_object* lean_float_to_string(double); 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_2607____closed__22; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instMonadAlwaysExceptStateT___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__10(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_object*, uint8_t); @@ -375,10 +378,9 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_printTraces___spec__ static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__14; LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_62_(lean_object*); 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_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__6___boxed(lean_object**); static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__20; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__4(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_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2607____closed__9; LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__8___boxed(lean_object**); 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; @@ -390,6 +392,7 @@ LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__2(lean_object*, l static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__14; 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___auto____x40_Lean_Util_Trace___hyg_2601____closed__5; static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__30; 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; @@ -400,13 +403,13 @@ lean_object* l_IO_println___at_Lean_instEval__1___spec__1(lean_object*, lean_obj 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_initFn____x40_Lean_Util_Trace___hyg_1421____closed__3; 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; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_printTraces___spec__4___rarg___lambda__2(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore(lean_object*); +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2601____closed__4; +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2601____closed__13; 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_Lean_trace_profiler_threshold_getSecs___closed__1; @@ -419,21 +422,22 @@ 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*); -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2607____closed__25; +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2601____closed__12; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_checkTraceOption_go___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__11(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_object*, lean_object*); static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__54; -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2607____closed__8; 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* l_Lean_KVMap_findCore(lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2601____closed__20; 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_initFn____x40_Lean_Util_Trace___hyg_1324_(lean_object*); LEAN_EXPORT lean_object* l_Lean_trace_profiler; static size_t l_Lean_instInhabitedTraceState___closed__3; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); @@ -449,7 +453,6 @@ 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_2607____closed__10; lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_instExceptToEmojiBool(lean_object*); static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__11; @@ -462,28 +465,25 @@ LEAN_EXPORT lean_object* l_Lean_withTraceNode(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_setTraceState(lean_object*); static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__16; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_printTraces___spec__3(lean_object*); +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2601____closed__9; LEAN_EXPORT lean_object* l_Lean_printTraces___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1421____closed__4; static lean_object* l_Lean_isTracingEnabledFor___rarg___closed__1; -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2607____closed__21; 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_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; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___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_object*); -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2607____closed__2; +static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1415____closed__1; 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_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_1330____closed__6; 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_initFn____x40_Lean_Util_Trace___hyg_1415____closed__5; static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__17; lean_object* l_StateRefT_x27_instMonadExceptOfStateRefT_x27___rarg(lean_object*); 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_Lean_initFn____x40_Lean_Util_Trace___hyg_1421____closed__5; -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; @@ -491,7 +491,6 @@ LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__11(lean_obj lean_object* l_Lean_MessageData_format(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instMonadAlwaysExceptReaderT(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__45; -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2607____closed__23; LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_withTraceNode_x27___spec__2___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__12; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___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*); @@ -499,8 +498,7 @@ static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doEle 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_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___spec__1___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2607____closed__11; +static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1415____closed__4; 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*); static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__26; @@ -2371,7 +2369,7 @@ x_10 = l_Lean_traceM___rarg___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_9); return x_10; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___spec__1(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNode___spec__1(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -2398,95 +2396,57 @@ goto _start; } } } -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___rarg___lambda__1(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_7 = l_Lean_PersistentArray_toArray___rarg(x_6); -x_8 = lean_array_get_size(x_7); -x_9 = lean_usize_of_nat(x_8); -lean_dec(x_8); -x_10 = 0; -x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___spec__1(x_9, x_10, x_7); -x_12 = lean_alloc_ctor(9, 3, 1); -lean_ctor_set(x_12, 0, x_1); -lean_ctor_set(x_12, 1, x_2); -lean_ctor_set(x_12, 2, x_11); -lean_ctor_set_uint8(x_12, sizeof(void*)*3, x_3); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_4); -lean_ctor_set(x_13, 1, x_12); -x_14 = l_Lean_PersistentArray_push___rarg(x_5, x_13); -return x_14; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, uint8_t x_6) { -_start: -{ -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_7 = lean_ctor_get(x_1, 0); -lean_inc(x_7); -lean_dec(x_1); -x_8 = lean_box(x_6); -x_9 = lean_alloc_closure((void*)(l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___rarg___lambda__1___boxed), 6, 5); -lean_closure_set(x_9, 0, x_3); -lean_closure_set(x_9, 1, x_5); -lean_closure_set(x_9, 2, x_8); -lean_closure_set(x_9, 3, x_4); -lean_closure_set(x_9, 4, x_2); -x_10 = lean_apply_1(x_7, x_9); -return x_10; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___rarg___boxed), 6, 0); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -size_t x_4; size_t x_5; lean_object* x_6; -x_4 = lean_unbox_usize(x_1); -lean_dec(x_1); -x_5 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___spec__1(x_4, x_5, x_3); +lean_object* x_5; lean_object* x_6; +x_5 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_5, 0, x_1); +lean_ctor_set(x_5, 1, x_2); +x_6 = l_Lean_PersistentArray_push___rarg(x_3, x_5); return x_6; } } -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -uint8_t x_7; lean_object* x_8; -x_7 = lean_unbox(x_3); -lean_dec(x_3); -x_8 = l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___rarg___lambda__1(x_1, x_2, x_7, x_4, x_5, x_6); -return x_8; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___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___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -uint8_t x_7; lean_object* x_8; -x_7 = lean_unbox(x_6); -lean_dec(x_6); -x_8 = l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___rarg(x_1, x_2, x_3, x_4, x_5, x_7); -return x_8; +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +lean_dec(x_1); +x_6 = lean_alloc_closure((void*)(l___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg___lambda__1___boxed), 4, 3); +lean_closure_set(x_6, 0, x_2); +lean_closure_set(x_6, 1, x_4); +lean_closure_set(x_6, 2, x_3); +x_7 = lean_apply_1(x_5, x_6); +return x_7; } } -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg___lambda__1(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 lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg___lambda__3(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_7; -x_7 = l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___rarg(x_1, x_2, x_3, x_4, x_6, x_5); -return x_7; +lean_object* x_10; lean_object* x_11; size_t x_12; size_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_10 = l_Lean_PersistentArray_toArray___rarg(x_9); +x_11 = lean_array_get_size(x_10); +x_12 = lean_usize_of_nat(x_11); +lean_dec(x_11); +x_13 = 0; +x_14 = l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNode___spec__1(x_12, x_13, x_10); +x_15 = lean_alloc_ctor(9, 3, 1); +lean_ctor_set(x_15, 0, x_1); +lean_ctor_set(x_15, 1, x_2); +lean_ctor_set(x_15, 2, x_14); +lean_ctor_set_uint8(x_15, sizeof(void*)*3, x_3); +x_16 = lean_apply_1(x_4, x_15); +x_17 = lean_alloc_closure((void*)(l___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg___lambda__2), 4, 3); +lean_closure_set(x_17, 0, x_5); +lean_closure_set(x_17, 1, x_6); +lean_closure_set(x_17, 2, x_7); +x_18 = lean_apply_4(x_8, lean_box(0), lean_box(0), x_16, x_17); +return x_18; } } -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; lean_object* x_7; @@ -2501,29 +2461,37 @@ return x_7; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___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, uint8_t x_9) { _start: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; x_10 = lean_ctor_get(x_1, 1); lean_inc(x_10); -lean_dec(x_1); -x_11 = lean_apply_1(x_4, x_8); -x_12 = lean_box(x_9); +x_11 = lean_ctor_get(x_2, 1); +lean_inc(x_11); +x_12 = lean_alloc_closure((void*)(l_Lean_getTraces___rarg___lambda__1), 2, 1); +lean_closure_set(x_12, 0, x_1); +lean_inc(x_10); +x_13 = lean_apply_4(x_10, lean_box(0), lean_box(0), x_11, x_12); +x_14 = lean_box(x_9); +lean_inc(x_10); lean_inc(x_7); -x_13 = lean_alloc_closure((void*)(l___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg___lambda__1___boxed), 6, 5); -lean_closure_set(x_13, 0, x_2); -lean_closure_set(x_13, 1, x_5); -lean_closure_set(x_13, 2, x_6); -lean_closure_set(x_13, 3, x_7); -lean_closure_set(x_13, 4, x_12); +x_15 = lean_alloc_closure((void*)(l___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg___lambda__3___boxed), 9, 8); +lean_closure_set(x_15, 0, x_6); +lean_closure_set(x_15, 1, x_8); +lean_closure_set(x_15, 2, x_14); +lean_closure_set(x_15, 3, x_4); +lean_closure_set(x_15, 4, x_2); +lean_closure_set(x_15, 5, x_7); +lean_closure_set(x_15, 6, x_5); +lean_closure_set(x_15, 7, x_10); lean_inc(x_10); -x_14 = lean_apply_4(x_10, lean_box(0), lean_box(0), x_11, x_13); -x_15 = lean_ctor_get(x_3, 0); -lean_inc(x_15); -x_16 = lean_alloc_closure((void*)(l___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg___lambda__2___boxed), 4, 3); -lean_closure_set(x_16, 0, x_7); -lean_closure_set(x_16, 1, x_3); -lean_closure_set(x_16, 2, x_14); -x_17 = lean_apply_4(x_10, lean_box(0), lean_box(0), x_15, x_16); -return x_17; +x_16 = lean_apply_4(x_10, lean_box(0), lean_box(0), x_13, x_15); +x_17 = lean_ctor_get(x_3, 0); +lean_inc(x_17); +x_18 = lean_alloc_closure((void*)(l___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg___lambda__4___boxed), 4, 3); +lean_closure_set(x_18, 0, x_7); +lean_closure_set(x_18, 1, x_3); +lean_closure_set(x_18, 2, x_16); +x_19 = lean_apply_4(x_10, lean_box(0), lean_box(0), x_17, x_18); +return x_19; } } LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode(lean_object* x_1) { @@ -2534,21 +2502,42 @@ x_2 = lean_alloc_closure((void*)(l___private_Lean_Util_Trace_0__Lean_addTraceNod return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNode___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -uint8_t x_7; lean_object* x_8; -x_7 = lean_unbox(x_5); -lean_dec(x_5); -x_8 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg___lambda__1(x_1, x_2, x_3, x_4, x_7, x_6); -return x_8; +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNode___spec__1(x_4, x_5, x_3); +return x_6; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg___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_Util_Trace_0__Lean_addTraceNode___rarg___lambda__1(x_1, x_2, x_3, x_4); +lean_dec(x_4); +return x_5; } } -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___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) { +_start: +{ +uint8_t x_10; lean_object* x_11; +x_10 = lean_unbox(x_3); +lean_dec(x_3); +x_11 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg___lambda__3(x_1, x_2, x_10, x_4, x_5, x_6, x_7, x_8, x_9); +return x_11; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg___lambda__2(x_1, x_2, x_3, x_4); +x_5 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg___lambda__4(x_1, x_2, x_3, x_4); lean_dec(x_4); lean_dec(x_1); return x_5; @@ -2671,7 +2660,7 @@ lean_dec(x_2); return x_5; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__1() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__1() { _start: { lean_object* x_1; @@ -2679,17 +2668,17 @@ x_1 = lean_mk_string_from_bytes("profiler", 8); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__2() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____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_1330____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____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_Trace___hyg_1330____closed__3() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__3() { _start: { lean_object* x_1; @@ -2697,13 +2686,13 @@ x_1 = lean_mk_string_from_bytes("activate nested traces with execution time over return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__4() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____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_Trace___hyg_1330____closed__1; -x_3 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__3; +x_2 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__1; +x_3 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__3; x_4 = lean_box(x_1); x_5 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_5, 0, x_4); @@ -2712,7 +2701,7 @@ lean_ctor_set(x_5, 2, x_3); return x_5; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__5() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__5() { _start: { lean_object* x_1; @@ -2720,29 +2709,29 @@ x_1 = lean_mk_string_from_bytes("Lean", 4); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__6() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__6() { _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_Trace___hyg_1330____closed__5; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__5; x_2 = l___private_Lean_Util_Trace_0__Lean_checkTraceOption___closed__1; -x_3 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__1; +x_3 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____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_1330_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324_(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_1330____closed__2; -x_3 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__4; -x_4 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__6; +x_2 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__2; +x_3 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__4; +x_4 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____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_initFn____x40_Lean_Util_Trace___hyg_1421____closed__1() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1415____closed__1() { _start: { lean_object* x_1; @@ -2750,18 +2739,18 @@ x_1 = lean_mk_string_from_bytes("threshold", 9); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1421____closed__2() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1415____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_1330____closed__1; -x_3 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1421____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__1; +x_3 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1415____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_1421____closed__3() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1415____closed__3() { _start: { lean_object* x_1; @@ -2769,13 +2758,13 @@ x_1 = lean_mk_string_from_bytes("threshold in milliseconds, traces below thresho return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1421____closed__4() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1415____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(10u); -x_2 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__1; -x_3 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1421____closed__3; +x_2 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__1; +x_3 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1415____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); @@ -2783,25 +2772,25 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1421____closed__5() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1415____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_initFn____x40_Lean_Util_Trace___hyg_1330____closed__5; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__5; x_2 = l___private_Lean_Util_Trace_0__Lean_checkTraceOption___closed__1; -x_3 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__1; -x_4 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1421____closed__1; +x_3 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__1; +x_4 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1415____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_1421_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1415_(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_1421____closed__2; -x_3 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1421____closed__4; -x_4 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1421____closed__5; +x_2 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1415____closed__2; +x_3 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1415____closed__4; +x_4 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1415____closed__5; x_5 = l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_41____spec__1(x_2, x_3, x_4, x_1); return x_5; } @@ -4975,7 +4964,7 @@ x_13 = l_Lean_withTraceNode_x27___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x return x_13; } } -static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2607____closed__1() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2601____closed__1() { _start: { lean_object* x_1; @@ -4983,7 +4972,7 @@ x_1 = lean_mk_string_from_bytes("Parser", 6); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2607____closed__2() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2601____closed__2() { _start: { lean_object* x_1; @@ -4991,7 +4980,7 @@ x_1 = lean_mk_string_from_bytes("Tactic", 6); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2607____closed__3() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2601____closed__3() { _start: { lean_object* x_1; @@ -4999,19 +4988,19 @@ x_1 = lean_mk_string_from_bytes("tacticSeq", 9); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2607____closed__4() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2601____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_initFn____x40_Lean_Util_Trace___hyg_1330____closed__5; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__1; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__2; -x_4 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__3; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__5; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2601____closed__1; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2601____closed__2; +x_4 = l___auto____x40_Lean_Util_Trace___hyg_2601____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_2607____closed__5() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2601____closed__5() { _start: { lean_object* x_1; @@ -5019,19 +5008,19 @@ x_1 = lean_mk_string_from_bytes("tacticSeq1Indented", 18); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2607____closed__6() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2601____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_initFn____x40_Lean_Util_Trace___hyg_1330____closed__5; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__1; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__2; -x_4 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__5; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__5; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2601____closed__1; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2601____closed__2; +x_4 = l___auto____x40_Lean_Util_Trace___hyg_2601____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_2607____closed__7() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2601____closed__7() { _start: { lean_object* x_1; @@ -5039,17 +5028,17 @@ x_1 = lean_mk_string_from_bytes("null", 4); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2607____closed__8() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2601____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_2607____closed__7; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2601____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_2607____closed__9() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2601____closed__9() { _start: { lean_object* x_1; @@ -5057,41 +5046,41 @@ x_1 = lean_mk_string_from_bytes("exact", 5); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2607____closed__10() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2601____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_initFn____x40_Lean_Util_Trace___hyg_1330____closed__5; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__1; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__2; -x_4 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__9; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__5; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2601____closed__1; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2601____closed__2; +x_4 = l___auto____x40_Lean_Util_Trace___hyg_2601____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_2607____closed__11() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2601____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_2607____closed__9; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2601____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_2607____closed__12() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2601____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_2607____closed__11; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2601____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_2607____closed__13() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2601____closed__13() { _start: { lean_object* x_1; @@ -5099,7 +5088,7 @@ x_1 = lean_mk_string_from_bytes("Term", 4); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2607____closed__14() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2601____closed__14() { _start: { lean_object* x_1; @@ -5107,19 +5096,19 @@ x_1 = lean_mk_string_from_bytes("declName", 8); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2607____closed__15() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2601____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_initFn____x40_Lean_Util_Trace___hyg_1330____closed__5; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__1; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__13; -x_4 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__14; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__5; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2601____closed__1; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2601____closed__13; +x_4 = l___auto____x40_Lean_Util_Trace___hyg_2601____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_2607____closed__16() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2601____closed__16() { _start: { lean_object* x_1; @@ -5127,35 +5116,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_2607____closed__17() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2601____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_2607____closed__16; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2601____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_2607____closed__18() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2601____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_2607____closed__17; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2601____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_2607____closed__19() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2601____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_2607____closed__15; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__18; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2601____closed__15; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2601____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); @@ -5163,23 +5152,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2607____closed__20() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2601____closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__12; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__19; +x_1 = l___auto____x40_Lean_Util_Trace___hyg_2601____closed__12; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2601____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_2607____closed__21() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2601____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_2607____closed__10; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__20; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2601____closed__10; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2601____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); @@ -5187,23 +5176,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2607____closed__22() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2601____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_2607____closed__21; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2601____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_2607____closed__23() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2601____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_2607____closed__8; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__22; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2601____closed__8; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2601____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); @@ -5211,23 +5200,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2607____closed__24() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2601____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_2607____closed__23; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2601____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_2607____closed__25() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2601____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_2607____closed__6; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__24; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2601____closed__6; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2601____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); @@ -5235,23 +5224,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2607____closed__26() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2601____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_2607____closed__25; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2601____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_2607____closed__27() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2601____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_2607____closed__4; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__26; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2601____closed__4; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2601____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); @@ -5259,11 +5248,11 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2607_() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2601_() { _start: { lean_object* x_1; -x_1 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__27; +x_1 = l___auto____x40_Lean_Util_Trace___hyg_2601____closed__27; return x_1; } } @@ -5412,7 +5401,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_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__5; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__5; x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; @@ -5668,9 +5657,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_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__5; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__1; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__13; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__5; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2601____closed__1; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2601____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; @@ -5696,9 +5685,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_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__5; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__1; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__13; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__5; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2601____closed__1; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2601____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; @@ -5716,9 +5705,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_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__5; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__1; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__13; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__5; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2601____closed__1; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2601____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; @@ -5736,9 +5725,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_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__5; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__1; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__13; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__5; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2601____closed__1; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2601____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; @@ -5764,9 +5753,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_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__5; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__1; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__13; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__5; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2601____closed__1; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2601____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; @@ -5784,9 +5773,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_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__5; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__1; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__13; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__5; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2601____closed__1; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2601____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; @@ -5839,9 +5828,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_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__5; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__1; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__13; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__5; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2601____closed__1; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2601____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; @@ -5867,9 +5856,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_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__5; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__1; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__13; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__5; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2601____closed__1; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2601____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; @@ -5887,9 +5876,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_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__5; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__1; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__13; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__5; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2601____closed__1; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2601____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; @@ -5915,9 +5904,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_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__5; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__1; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__13; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__5; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2601____closed__1; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2601____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; @@ -5943,9 +5932,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_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__5; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__1; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__13; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__5; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2601____closed__1; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2601____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; @@ -5980,7 +5969,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_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__5; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____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; @@ -6060,9 +6049,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_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__5; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__1; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__13; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__5; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2601____closed__1; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2601____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; @@ -6097,7 +6086,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_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__5; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____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; @@ -6161,9 +6150,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_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__5; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__1; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__13; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__5; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2601____closed__1; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2601____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; @@ -6217,7 +6206,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_2607____closed__8; +x_14 = l___auto____x40_Lean_Util_Trace___hyg_2601____closed__8; x_15 = l_Lean_instInhabitedTraceState___closed__1; lean_inc(x_7); x_16 = lean_alloc_ctor(1, 3, 0); @@ -6437,9 +6426,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_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__5; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__1; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_2607____closed__13; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__5; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2601____closed__1; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2601____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; @@ -6484,7 +6473,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_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__5; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____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; @@ -6548,7 +6537,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_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__5; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____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; @@ -6640,7 +6629,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_2607____closed__8; +x_30 = l___auto____x40_Lean_Util_Trace___hyg_2601____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; @@ -7017,17 +7006,18 @@ x_5 = l_MonadExcept_ofExcept___at_Lean_withTraceNodeBefore___spec__1___rarg(x_1, return x_5; } } -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, 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_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, lean_object* x_13) { _start: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___rarg(x_1, x_2, x_3, x_4, x_10, x_5); -x_13 = lean_alloc_closure((void*)(l_Lean_withTraceNodeBefore___rarg___lambda__1), 4, 3); -lean_closure_set(x_13, 0, x_6); -lean_closure_set(x_13, 1, x_7); -lean_closure_set(x_13, 2, x_8); -x_14 = lean_apply_4(x_9, lean_box(0), lean_box(0), x_12, x_13); -return x_14; +lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_inc(x_1); +x_14 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_12, x_8); +x_15 = lean_alloc_closure((void*)(l_Lean_withTraceNodeBefore___rarg___lambda__1), 4, 3); +lean_closure_set(x_15, 0, x_1); +lean_closure_set(x_15, 1, x_9); +lean_closure_set(x_15, 2, x_10); +x_16 = lean_apply_4(x_11, lean_box(0), lean_box(0), x_14, x_15); +return x_16; } } static lean_object* _init_l_Lean_withTraceNodeBefore___rarg___lambda__3___closed__1() { @@ -7047,148 +7037,154 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -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, 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) { +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, uint8_t x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, uint8_t x_15, double x_16, lean_object* x_17) { _start: { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t 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; lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; 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_18); -lean_ctor_set(x_19, 1, x_17); -x_20 = l_Lean_withTraceNodeBefore___rarg___lambda__3___closed__2; +x_18 = lean_apply_1(x_1, x_2); +x_19 = l_Lean_stringToMessageData(x_18); +lean_dec(x_18); +x_20 = l_Lean_withTraceNode___rarg___lambda__5___closed__6; 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_3); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_19); +x_22 = l_Lean_withTraceNodeBefore___rarg___lambda__3___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_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_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +x_24 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_3); +x_25 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_20); +x_26 = l_Lean_shouldProfile___rarg___lambda__1___closed__1; +x_27 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_14, x_26); +if (x_27 == 0) { -if (x_13 == 0) +if (x_15 == 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__2___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; +lean_object* x_28; 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_28 = lean_ctor_get(x_4, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_28, 1); +lean_inc(x_29); +lean_dec(x_28); +x_30 = lean_box(0); +x_31 = lean_apply_2(x_29, lean_box(0), x_30); +x_32 = lean_box(x_11); +lean_inc(x_13); +x_33 = lean_alloc_closure((void*)(l_Lean_withTraceNodeBefore___rarg___lambda__2___boxed), 13, 12); +lean_closure_set(x_33, 0, x_4); +lean_closure_set(x_33, 1, x_5); +lean_closure_set(x_33, 2, x_6); +lean_closure_set(x_33, 3, x_7); +lean_closure_set(x_33, 4, x_8); +lean_closure_set(x_33, 5, x_9); +lean_closure_set(x_33, 6, x_10); +lean_closure_set(x_33, 7, x_32); +lean_closure_set(x_33, 8, x_12); +lean_closure_set(x_33, 9, x_2); +lean_closure_set(x_33, 10, x_13); +lean_closure_set(x_33, 11, x_25); +x_34 = lean_apply_4(x_13, lean_box(0), lean_box(0), x_31, x_33); +return x_34; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; -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); -x_36 = l_Lean_withTraceNode___rarg___lambda__5___closed__2; -x_37 = lean_alloc_ctor(7, 2, 0); +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_35 = lean_float_to_string(x_16); +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(x_37, 1, x_35); -x_38 = l_Lean_withTraceNode___rarg___lambda__5___closed__4; +x_38 = l_Lean_withTraceNode___rarg___lambda__5___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); -x_40 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_40, 0, x_39); -lean_ctor_set(x_40, 1, x_23); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_37); +x_40 = l_Lean_withTraceNode___rarg___lambda__5___closed__4; 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__2___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; +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_25); +x_43 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_20); +x_44 = lean_ctor_get(x_4, 0); +lean_inc(x_44); +x_45 = lean_ctor_get(x_44, 1); +lean_inc(x_45); +lean_dec(x_44); +x_46 = lean_box(0); +x_47 = lean_apply_2(x_45, lean_box(0), x_46); +x_48 = lean_box(x_11); +lean_inc(x_13); +x_49 = lean_alloc_closure((void*)(l_Lean_withTraceNodeBefore___rarg___lambda__2___boxed), 13, 12); +lean_closure_set(x_49, 0, x_4); +lean_closure_set(x_49, 1, x_5); +lean_closure_set(x_49, 2, x_6); +lean_closure_set(x_49, 3, x_7); +lean_closure_set(x_49, 4, x_8); +lean_closure_set(x_49, 5, x_9); +lean_closure_set(x_49, 6, x_10); +lean_closure_set(x_49, 7, x_48); +lean_closure_set(x_49, 8, x_12); +lean_closure_set(x_49, 9, x_2); +lean_closure_set(x_49, 10, x_13); +lean_closure_set(x_49, 11, x_43); +x_50 = lean_apply_4(x_13, lean_box(0), lean_box(0), x_47, x_49); +return x_50; } } 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_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_51 = lean_float_to_string(x_16); +x_52 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_52, 0, x_51); +x_53 = lean_alloc_ctor(0, 1, 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_54 = l_Lean_withTraceNode___rarg___lambda__5___closed__2; 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); +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_53); +x_56 = l_Lean_withTraceNode___rarg___lambda__5___closed__4; 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__2___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_ctor_set(x_57, 0, x_55); +lean_ctor_set(x_57, 1, x_56); +x_58 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_58, 0, x_57); +lean_ctor_set(x_58, 1, x_25); +x_59 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_59, 0, x_58); +lean_ctor_set(x_59, 1, x_20); +x_60 = lean_ctor_get(x_4, 0); +lean_inc(x_60); +x_61 = lean_ctor_get(x_60, 1); +lean_inc(x_61); +lean_dec(x_60); +x_62 = lean_box(0); +x_63 = lean_apply_2(x_61, lean_box(0), x_62); +x_64 = lean_box(x_11); +lean_inc(x_13); +x_65 = lean_alloc_closure((void*)(l_Lean_withTraceNodeBefore___rarg___lambda__2___boxed), 13, 12); +lean_closure_set(x_65, 0, x_4); +lean_closure_set(x_65, 1, x_5); +lean_closure_set(x_65, 2, x_6); +lean_closure_set(x_65, 3, x_7); +lean_closure_set(x_65, 4, x_8); +lean_closure_set(x_65, 5, x_9); +lean_closure_set(x_65, 6, x_10); +lean_closure_set(x_65, 7, x_64); +lean_closure_set(x_65, 8, x_12); +lean_closure_set(x_65, 9, x_2); +lean_closure_set(x_65, 10, x_13); +lean_closure_set(x_65, 11, x_59); +x_66 = lean_apply_4(x_13, lean_box(0), lean_box(0), x_63, x_65); +return x_66; } } } @@ -7204,156 +7200,160 @@ 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_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, 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) { +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, uint8_t 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_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) +lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_38; uint8_t x_39; +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_38 = l_Lean_shouldProfile___rarg___lambda__1___closed__2; +x_39 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_13, x_38); +if (x_39 == 0) { -uint8_t x_38; -x_38 = 0; -x_16 = x_38; -goto block_35; +uint8_t x_40; +x_40 = 0; +x_18 = x_40; +goto block_37; } 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; +double x_41; double x_42; uint8_t x_43; +x_41 = l_Lean_trace_profiler_threshold_getSecs(x_13); +x_42 = lean_unbox_float(x_17); +x_43 = lean_float_decLt(x_41, x_42); +x_18 = x_43; +goto block_37; } -block_35: +block_37: { -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_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_box(x_10); +x_20 = lean_box(x_18); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_7); lean_inc(x_4); lean_inc(x_3); -lean_inc(x_14); -x_19 = lean_alloc_closure((void*)(l_Lean_withTraceNodeBefore___rarg___lambda__3___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) +lean_inc(x_16); +x_21 = lean_alloc_closure((void*)(l_Lean_withTraceNodeBefore___rarg___lambda__3___boxed), 17, 16); +lean_closure_set(x_21, 0, x_1); +lean_closure_set(x_21, 1, x_16); +lean_closure_set(x_21, 2, x_2); +lean_closure_set(x_21, 3, x_3); +lean_closure_set(x_21, 4, x_4); +lean_closure_set(x_21, 5, x_5); +lean_closure_set(x_21, 6, x_6); +lean_closure_set(x_21, 7, x_7); +lean_closure_set(x_21, 8, x_8); +lean_closure_set(x_21, 9, x_9); +lean_closure_set(x_21, 10, x_19); +lean_closure_set(x_21, 11, x_11); +lean_closure_set(x_21, 12, x_12); +lean_closure_set(x_21, 13, x_13); +lean_closure_set(x_21, 14, x_20); +lean_closure_set(x_21, 15, x_17); +if (x_14 == 0) { -if (x_16 == 0) +if (x_18 == 0) { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -lean_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_withTraceNodeBefore___rarg___lambda__4___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; +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_dec(x_21); +x_22 = lean_ctor_get(x_4, 0); +lean_inc(x_22); +lean_dec(x_4); +x_23 = lean_alloc_closure((void*)(l_Lean_PersistentArray_append___rarg), 2, 1); +lean_closure_set(x_23, 0, x_7); +x_24 = lean_apply_1(x_22, x_23); +lean_inc(x_12); +x_25 = lean_alloc_closure((void*)(l_Lean_withTraceNodeBefore___rarg___lambda__4___boxed), 5, 4); +lean_closure_set(x_25, 0, x_3); +lean_closure_set(x_25, 1, x_11); +lean_closure_set(x_25, 2, x_16); +lean_closure_set(x_25, 3, x_12); +x_26 = lean_apply_4(x_12, lean_box(0), lean_box(0), x_24, x_25); +return x_26; } 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_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +lean_dec(x_16); +lean_dec(x_11); +lean_dec(x_7); lean_dec(x_4); +x_27 = lean_ctor_get(x_3, 0); +lean_inc(x_27); 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; +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_12, lean_box(0), lean_box(0), x_30, x_21); +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_dec(x_14); -lean_dec(x_9); +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +lean_dec(x_16); +lean_dec(x_11); +lean_dec(x_7); lean_dec(x_4); +x_32 = lean_ctor_get(x_3, 0); +lean_inc(x_32); 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; +x_33 = lean_ctor_get(x_32, 1); +lean_inc(x_33); +lean_dec(x_32); +x_34 = lean_box(0); +x_35 = lean_apply_2(x_33, lean_box(0), x_34); +x_36 = lean_apply_4(x_12, lean_box(0), lean_box(0), x_35, x_21); +return x_36; } } } } -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, uint8_t x_11, lean_object* x_12, lean_object* x_13, uint8_t x_14, lean_object* x_15) { +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, lean_object* x_14, lean_object* x_15, uint8_t x_16, lean_object* x_17) { _start: { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +lean_object* x_18; lean_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_inc(x_1); -x_16 = lean_alloc_closure((void*)(l_Lean_withTraceNode___rarg___lambda__1), 2, 1); -lean_closure_set(x_16, 0, x_1); +x_18 = lean_alloc_closure((void*)(l_Lean_withTraceNode___rarg___lambda__1), 2, 1); +lean_closure_set(x_18, 0, x_1); lean_inc(x_2); -x_17 = lean_apply_4(x_2, lean_box(0), lean_box(0), x_3, x_16); +x_19 = lean_apply_4(x_2, lean_box(0), lean_box(0), x_3, x_18); lean_inc(x_1); -x_18 = lean_alloc_closure((void*)(l_Lean_withTraceNode___rarg___lambda__2), 2, 1); -lean_closure_set(x_18, 0, x_1); -x_19 = lean_ctor_get(x_4, 1); -lean_inc(x_19); +x_20 = lean_alloc_closure((void*)(l_Lean_withTraceNode___rarg___lambda__2), 2, 1); +lean_closure_set(x_20, 0, x_1); +x_21 = lean_ctor_get(x_4, 1); +lean_inc(x_21); lean_dec(x_4); -x_20 = lean_apply_3(x_19, lean_box(0), x_17, x_18); +x_22 = lean_apply_3(x_21, lean_box(0), x_19, x_20); lean_inc(x_1); -x_21 = l_Lean_withSeconds___rarg(x_1, x_5, x_20); -x_22 = lean_box(x_11); -x_23 = lean_box(x_14); +x_23 = l_Lean_withSeconds___rarg(x_1, x_5, x_22); +x_24 = lean_box(x_13); +x_25 = lean_box(x_16); lean_inc(x_2); -x_24 = lean_alloc_closure((void*)(l_Lean_withTraceNodeBefore___rarg___lambda__5___boxed), 13, 12); -lean_closure_set(x_24, 0, x_6); -lean_closure_set(x_24, 1, x_15); -lean_closure_set(x_24, 2, x_7); -lean_closure_set(x_24, 3, x_8); -lean_closure_set(x_24, 4, x_9); -lean_closure_set(x_24, 5, x_10); -lean_closure_set(x_24, 6, x_22); -lean_closure_set(x_24, 7, x_1); -lean_closure_set(x_24, 8, x_12); -lean_closure_set(x_24, 9, x_2); -lean_closure_set(x_24, 10, x_13); -lean_closure_set(x_24, 11, x_23); -x_25 = lean_apply_4(x_2, lean_box(0), lean_box(0), x_21, x_24); -return x_25; +x_26 = lean_alloc_closure((void*)(l_Lean_withTraceNodeBefore___rarg___lambda__5___boxed), 15, 14); +lean_closure_set(x_26, 0, x_6); +lean_closure_set(x_26, 1, x_17); +lean_closure_set(x_26, 2, x_1); +lean_closure_set(x_26, 3, x_7); +lean_closure_set(x_26, 4, x_8); +lean_closure_set(x_26, 5, x_9); +lean_closure_set(x_26, 6, x_10); +lean_closure_set(x_26, 7, x_11); +lean_closure_set(x_26, 8, x_12); +lean_closure_set(x_26, 9, x_24); +lean_closure_set(x_26, 10, x_14); +lean_closure_set(x_26, 11, x_2); +lean_closure_set(x_26, 12, x_15); +lean_closure_set(x_26, 13, x_25); +x_27 = lean_apply_4(x_2, lean_box(0), lean_box(0), x_23, x_26); +return x_27; } } 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, lean_object* x_12, lean_object* x_13, uint8_t x_14, lean_object* x_15, lean_object* x_16, uint8_t x_17, lean_object* x_18) { @@ -7361,9 +7361,11 @@ LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__7(lean_obje { lean_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_inc(x_2); +lean_inc(x_1); x_19 = lean_apply_4(x_2, lean_box(0), lean_box(0), x_3, x_1); +lean_inc(x_4); lean_inc(x_18); -x_20 = lean_alloc_closure((void*)(l___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg___lambda__2___boxed), 4, 3); +x_20 = lean_alloc_closure((void*)(l___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg___lambda__4___boxed), 4, 3); lean_closure_set(x_20, 0, x_18); lean_closure_set(x_20, 1, x_4); lean_closure_set(x_20, 2, x_19); @@ -7372,7 +7374,7 @@ x_21 = lean_apply_4(x_2, lean_box(0), lean_box(0), x_5, x_20); x_22 = lean_box(x_14); x_23 = lean_box(x_17); lean_inc(x_2); -x_24 = lean_alloc_closure((void*)(l_Lean_withTraceNodeBefore___rarg___lambda__6___boxed), 15, 14); +x_24 = lean_alloc_closure((void*)(l_Lean_withTraceNodeBefore___rarg___lambda__6___boxed), 17, 16); lean_closure_set(x_24, 0, x_6); lean_closure_set(x_24, 1, x_2); lean_closure_set(x_24, 2, x_7); @@ -7380,13 +7382,15 @@ lean_closure_set(x_24, 3, x_8); lean_closure_set(x_24, 4, x_9); lean_closure_set(x_24, 5, x_10); lean_closure_set(x_24, 6, x_11); -lean_closure_set(x_24, 7, x_12); -lean_closure_set(x_24, 8, x_13); -lean_closure_set(x_24, 9, x_18); -lean_closure_set(x_24, 10, x_22); -lean_closure_set(x_24, 11, x_15); -lean_closure_set(x_24, 12, x_16); -lean_closure_set(x_24, 13, x_23); +lean_closure_set(x_24, 7, x_4); +lean_closure_set(x_24, 8, x_1); +lean_closure_set(x_24, 9, x_12); +lean_closure_set(x_24, 10, x_13); +lean_closure_set(x_24, 11, x_18); +lean_closure_set(x_24, 12, x_22); +lean_closure_set(x_24, 13, x_15); +lean_closure_set(x_24, 14, x_16); +lean_closure_set(x_24, 15, x_23); x_25 = lean_apply_4(x_2, lean_box(0), lean_box(0), x_21, x_24); return x_25; } @@ -7594,31 +7598,48 @@ x_3 = lean_alloc_closure((void*)(l_Lean_withTraceNodeBefore___rarg___boxed), 14, return x_3; } } -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) { _start: { -uint8_t x_12; lean_object* x_13; -x_12 = lean_unbox(x_5); -lean_dec(x_5); -x_13 = l_Lean_withTraceNodeBefore___rarg___lambda__2(x_1, x_2, x_3, x_4, x_12, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_11); -return x_13; +uint8_t x_14; lean_object* x_15; +x_14 = lean_unbox(x_8); +lean_dec(x_8); +x_15 = l_Lean_withTraceNodeBefore___rarg___lambda__2(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); +lean_dec(x_13); +return x_15; } } -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, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___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]; _start: { -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__3(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); +uint8_t x_18; uint8_t x_19; double x_20; lean_object* x_21; +x_18 = lean_unbox(x_11); +lean_dec(x_11); +x_19 = lean_unbox(x_15); lean_dec(x_15); -lean_dec(x_12); -return x_19; +x_20 = lean_unbox_float(x_16); +lean_dec(x_16); +x_21 = 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_18, x_12, x_13, x_14, x_19, x_20, x_17); +lean_dec(x_17); +lean_dec(x_14); +return x_21; } } 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) { @@ -7630,30 +7651,47 @@ lean_dec(x_5); return x_6; } } -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) { -_start: -{ -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_16 = l_Lean_withTraceNodeBefore___rarg___lambda__5(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__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_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) { _start: { uint8_t x_16; uint8_t x_17; lean_object* x_18; -x_16 = lean_unbox(x_11); -lean_dec(x_11); +x_16 = lean_unbox(x_10); +lean_dec(x_10); 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_16, x_12, x_13, x_17, x_15); +x_18 = 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_16, x_11, x_12, x_13, x_17, x_15); return x_18; } } +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__6___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +_start: +{ +uint8_t x_18; uint8_t x_19; lean_object* x_20; +x_18 = lean_unbox(x_13); +lean_dec(x_13); +x_19 = lean_unbox(x_16); +lean_dec(x_16); +x_20 = 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_18, x_14, x_15, x_19, x_17); +return x_20; +} +} LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__7___boxed(lean_object** _args) { lean_object* x_1 = _args[0]; lean_object* x_2 = _args[1]; @@ -7795,34 +7833,34 @@ 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_initFn____x40_Lean_Util_Trace___hyg_1330____closed__1 = _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__1); -l_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__2 = _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__2(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__2); -l_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__3 = _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__3(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__3); -l_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__4 = _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__4(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__4); -l_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__5 = _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__5(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__5); -l_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__6 = _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__6(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Trace___hyg_1330____closed__6); -if (builtin) {res = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1330_(lean_io_mk_world()); +l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__1 = _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__1); +l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__2 = _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__2); +l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__3 = _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__3); +l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__4 = _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__4(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__4); +l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__5 = _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__5(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__5); +l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__6 = _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__6(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324____closed__6); +if (builtin) {res = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1324_(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_1421____closed__1 = _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1421____closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Trace___hyg_1421____closed__1); -l_Lean_initFn____x40_Lean_Util_Trace___hyg_1421____closed__2 = _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1421____closed__2(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Trace___hyg_1421____closed__2); -l_Lean_initFn____x40_Lean_Util_Trace___hyg_1421____closed__3 = _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1421____closed__3(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Trace___hyg_1421____closed__3); -l_Lean_initFn____x40_Lean_Util_Trace___hyg_1421____closed__4 = _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1421____closed__4(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Trace___hyg_1421____closed__4); -l_Lean_initFn____x40_Lean_Util_Trace___hyg_1421____closed__5 = _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1421____closed__5(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Trace___hyg_1421____closed__5); -if (builtin) {res = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1421_(lean_io_mk_world()); +}l_Lean_initFn____x40_Lean_Util_Trace___hyg_1415____closed__1 = _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1415____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Trace___hyg_1415____closed__1); +l_Lean_initFn____x40_Lean_Util_Trace___hyg_1415____closed__2 = _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1415____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Trace___hyg_1415____closed__2); +l_Lean_initFn____x40_Lean_Util_Trace___hyg_1415____closed__3 = _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1415____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Trace___hyg_1415____closed__3); +l_Lean_initFn____x40_Lean_Util_Trace___hyg_1415____closed__4 = _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1415____closed__4(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Trace___hyg_1415____closed__4); +l_Lean_initFn____x40_Lean_Util_Trace___hyg_1415____closed__5 = _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1415____closed__5(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Trace___hyg_1415____closed__5); +if (builtin) {res = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1415_(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); @@ -7854,62 +7892,62 @@ l_Lean_withTraceNode___rarg___lambda__6___closed__2 = _init_l_Lean_withTraceNode lean_mark_persistent(l_Lean_withTraceNode___rarg___lambda__6___closed__2); 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_2607____closed__1 = _init_l___auto____x40_Lean_Util_Trace___hyg_2607____closed__1(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2607____closed__1); -l___auto____x40_Lean_Util_Trace___hyg_2607____closed__2 = _init_l___auto____x40_Lean_Util_Trace___hyg_2607____closed__2(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2607____closed__2); -l___auto____x40_Lean_Util_Trace___hyg_2607____closed__3 = _init_l___auto____x40_Lean_Util_Trace___hyg_2607____closed__3(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2607____closed__3); -l___auto____x40_Lean_Util_Trace___hyg_2607____closed__4 = _init_l___auto____x40_Lean_Util_Trace___hyg_2607____closed__4(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2607____closed__4); -l___auto____x40_Lean_Util_Trace___hyg_2607____closed__5 = _init_l___auto____x40_Lean_Util_Trace___hyg_2607____closed__5(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2607____closed__5); -l___auto____x40_Lean_Util_Trace___hyg_2607____closed__6 = _init_l___auto____x40_Lean_Util_Trace___hyg_2607____closed__6(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2607____closed__6); -l___auto____x40_Lean_Util_Trace___hyg_2607____closed__7 = _init_l___auto____x40_Lean_Util_Trace___hyg_2607____closed__7(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2607____closed__7); -l___auto____x40_Lean_Util_Trace___hyg_2607____closed__8 = _init_l___auto____x40_Lean_Util_Trace___hyg_2607____closed__8(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2607____closed__8); -l___auto____x40_Lean_Util_Trace___hyg_2607____closed__9 = _init_l___auto____x40_Lean_Util_Trace___hyg_2607____closed__9(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2607____closed__9); -l___auto____x40_Lean_Util_Trace___hyg_2607____closed__10 = _init_l___auto____x40_Lean_Util_Trace___hyg_2607____closed__10(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2607____closed__10); -l___auto____x40_Lean_Util_Trace___hyg_2607____closed__11 = _init_l___auto____x40_Lean_Util_Trace___hyg_2607____closed__11(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2607____closed__11); -l___auto____x40_Lean_Util_Trace___hyg_2607____closed__12 = _init_l___auto____x40_Lean_Util_Trace___hyg_2607____closed__12(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2607____closed__12); -l___auto____x40_Lean_Util_Trace___hyg_2607____closed__13 = _init_l___auto____x40_Lean_Util_Trace___hyg_2607____closed__13(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2607____closed__13); -l___auto____x40_Lean_Util_Trace___hyg_2607____closed__14 = _init_l___auto____x40_Lean_Util_Trace___hyg_2607____closed__14(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2607____closed__14); -l___auto____x40_Lean_Util_Trace___hyg_2607____closed__15 = _init_l___auto____x40_Lean_Util_Trace___hyg_2607____closed__15(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2607____closed__15); -l___auto____x40_Lean_Util_Trace___hyg_2607____closed__16 = _init_l___auto____x40_Lean_Util_Trace___hyg_2607____closed__16(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2607____closed__16); -l___auto____x40_Lean_Util_Trace___hyg_2607____closed__17 = _init_l___auto____x40_Lean_Util_Trace___hyg_2607____closed__17(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2607____closed__17); -l___auto____x40_Lean_Util_Trace___hyg_2607____closed__18 = _init_l___auto____x40_Lean_Util_Trace___hyg_2607____closed__18(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2607____closed__18); -l___auto____x40_Lean_Util_Trace___hyg_2607____closed__19 = _init_l___auto____x40_Lean_Util_Trace___hyg_2607____closed__19(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2607____closed__19); -l___auto____x40_Lean_Util_Trace___hyg_2607____closed__20 = _init_l___auto____x40_Lean_Util_Trace___hyg_2607____closed__20(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2607____closed__20); -l___auto____x40_Lean_Util_Trace___hyg_2607____closed__21 = _init_l___auto____x40_Lean_Util_Trace___hyg_2607____closed__21(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2607____closed__21); -l___auto____x40_Lean_Util_Trace___hyg_2607____closed__22 = _init_l___auto____x40_Lean_Util_Trace___hyg_2607____closed__22(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2607____closed__22); -l___auto____x40_Lean_Util_Trace___hyg_2607____closed__23 = _init_l___auto____x40_Lean_Util_Trace___hyg_2607____closed__23(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2607____closed__23); -l___auto____x40_Lean_Util_Trace___hyg_2607____closed__24 = _init_l___auto____x40_Lean_Util_Trace___hyg_2607____closed__24(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2607____closed__24); -l___auto____x40_Lean_Util_Trace___hyg_2607____closed__25 = _init_l___auto____x40_Lean_Util_Trace___hyg_2607____closed__25(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2607____closed__25); -l___auto____x40_Lean_Util_Trace___hyg_2607____closed__26 = _init_l___auto____x40_Lean_Util_Trace___hyg_2607____closed__26(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2607____closed__26); -l___auto____x40_Lean_Util_Trace___hyg_2607____closed__27 = _init_l___auto____x40_Lean_Util_Trace___hyg_2607____closed__27(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2607____closed__27); -l___auto____x40_Lean_Util_Trace___hyg_2607_ = _init_l___auto____x40_Lean_Util_Trace___hyg_2607_(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2607_); +l___auto____x40_Lean_Util_Trace___hyg_2601____closed__1 = _init_l___auto____x40_Lean_Util_Trace___hyg_2601____closed__1(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2601____closed__1); +l___auto____x40_Lean_Util_Trace___hyg_2601____closed__2 = _init_l___auto____x40_Lean_Util_Trace___hyg_2601____closed__2(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2601____closed__2); +l___auto____x40_Lean_Util_Trace___hyg_2601____closed__3 = _init_l___auto____x40_Lean_Util_Trace___hyg_2601____closed__3(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2601____closed__3); +l___auto____x40_Lean_Util_Trace___hyg_2601____closed__4 = _init_l___auto____x40_Lean_Util_Trace___hyg_2601____closed__4(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2601____closed__4); +l___auto____x40_Lean_Util_Trace___hyg_2601____closed__5 = _init_l___auto____x40_Lean_Util_Trace___hyg_2601____closed__5(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2601____closed__5); +l___auto____x40_Lean_Util_Trace___hyg_2601____closed__6 = _init_l___auto____x40_Lean_Util_Trace___hyg_2601____closed__6(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2601____closed__6); +l___auto____x40_Lean_Util_Trace___hyg_2601____closed__7 = _init_l___auto____x40_Lean_Util_Trace___hyg_2601____closed__7(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2601____closed__7); +l___auto____x40_Lean_Util_Trace___hyg_2601____closed__8 = _init_l___auto____x40_Lean_Util_Trace___hyg_2601____closed__8(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2601____closed__8); +l___auto____x40_Lean_Util_Trace___hyg_2601____closed__9 = _init_l___auto____x40_Lean_Util_Trace___hyg_2601____closed__9(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2601____closed__9); +l___auto____x40_Lean_Util_Trace___hyg_2601____closed__10 = _init_l___auto____x40_Lean_Util_Trace___hyg_2601____closed__10(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2601____closed__10); +l___auto____x40_Lean_Util_Trace___hyg_2601____closed__11 = _init_l___auto____x40_Lean_Util_Trace___hyg_2601____closed__11(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2601____closed__11); +l___auto____x40_Lean_Util_Trace___hyg_2601____closed__12 = _init_l___auto____x40_Lean_Util_Trace___hyg_2601____closed__12(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2601____closed__12); +l___auto____x40_Lean_Util_Trace___hyg_2601____closed__13 = _init_l___auto____x40_Lean_Util_Trace___hyg_2601____closed__13(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2601____closed__13); +l___auto____x40_Lean_Util_Trace___hyg_2601____closed__14 = _init_l___auto____x40_Lean_Util_Trace___hyg_2601____closed__14(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2601____closed__14); +l___auto____x40_Lean_Util_Trace___hyg_2601____closed__15 = _init_l___auto____x40_Lean_Util_Trace___hyg_2601____closed__15(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2601____closed__15); +l___auto____x40_Lean_Util_Trace___hyg_2601____closed__16 = _init_l___auto____x40_Lean_Util_Trace___hyg_2601____closed__16(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2601____closed__16); +l___auto____x40_Lean_Util_Trace___hyg_2601____closed__17 = _init_l___auto____x40_Lean_Util_Trace___hyg_2601____closed__17(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2601____closed__17); +l___auto____x40_Lean_Util_Trace___hyg_2601____closed__18 = _init_l___auto____x40_Lean_Util_Trace___hyg_2601____closed__18(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2601____closed__18); +l___auto____x40_Lean_Util_Trace___hyg_2601____closed__19 = _init_l___auto____x40_Lean_Util_Trace___hyg_2601____closed__19(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2601____closed__19); +l___auto____x40_Lean_Util_Trace___hyg_2601____closed__20 = _init_l___auto____x40_Lean_Util_Trace___hyg_2601____closed__20(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2601____closed__20); +l___auto____x40_Lean_Util_Trace___hyg_2601____closed__21 = _init_l___auto____x40_Lean_Util_Trace___hyg_2601____closed__21(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2601____closed__21); +l___auto____x40_Lean_Util_Trace___hyg_2601____closed__22 = _init_l___auto____x40_Lean_Util_Trace___hyg_2601____closed__22(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2601____closed__22); +l___auto____x40_Lean_Util_Trace___hyg_2601____closed__23 = _init_l___auto____x40_Lean_Util_Trace___hyg_2601____closed__23(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2601____closed__23); +l___auto____x40_Lean_Util_Trace___hyg_2601____closed__24 = _init_l___auto____x40_Lean_Util_Trace___hyg_2601____closed__24(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2601____closed__24); +l___auto____x40_Lean_Util_Trace___hyg_2601____closed__25 = _init_l___auto____x40_Lean_Util_Trace___hyg_2601____closed__25(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2601____closed__25); +l___auto____x40_Lean_Util_Trace___hyg_2601____closed__26 = _init_l___auto____x40_Lean_Util_Trace___hyg_2601____closed__26(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2601____closed__26); +l___auto____x40_Lean_Util_Trace___hyg_2601____closed__27 = _init_l___auto____x40_Lean_Util_Trace___hyg_2601____closed__27(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2601____closed__27); +l___auto____x40_Lean_Util_Trace___hyg_2601_ = _init_l___auto____x40_Lean_Util_Trace___hyg_2601_(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2601_); 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/InteractiveDiagnostic.c b/stage0/stdlib/Lean/Widget/InteractiveDiagnostic.c index 0061abe76afd..db03604c7099 100644 --- a/stage0/stdlib/Lean/Widget/InteractiveDiagnostic.c +++ b/stage0/stdlib/Lean/Widget/InteractiveDiagnostic.c @@ -167,7 +167,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_W static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_Lean_Widget_MsgEmbed_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_708____lambda__1___closed__18; lean_object* lean_nat_to_int(lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_Lean_Widget_StrictOrLazy_instRpcEncodableStrictOrLazy_enc____x40_Lean_Widget_InteractiveDiagnostic___hyg_4_(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_dbgTraceVal___at___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_chopUpChildren___spec__1___lambda__1(lean_object*, lean_object*); lean_object* l_Subarray_size___rarg(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_Lean_Lsp_DiagnosticWith_instRpcEncodableDiagnosticWith_enc____x40_Lean_Widget_InteractiveDiagnostic___hyg_1549____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_instTypeNameLazyTraceChildren; @@ -281,7 +280,6 @@ static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget LEAN_EXPORT lean_object* l_Lean_Option_get_x3f___at___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_go___spec__6(lean_object*, lean_object*); static lean_object* l_Lean_Widget_instInhabitedEmbedFmt___closed__2; lean_object* l_Lean_Json_parseTagged(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_dbgTraceVal___at___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_chopUpChildren___spec__1___lambda__1___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_mkContextInfo___closed__3; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_go___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_string_length(lean_object*); @@ -309,7 +307,6 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_msgToInteractiv lean_object* l_String_toName(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_Lean_Widget_StrictOrLazy_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_89_(lean_object*); static lean_object* l_Lean_Widget_Lean_Widget_StrictOrLazy_instFromJsonRpcEncodablePacket___closed__1; -LEAN_EXPORT lean_object* l_dbgTraceVal___at___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_chopUpChildren___spec__1(lean_object*); static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_Lean_Lsp_DiagnosticWith_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1634____closed__10; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_Lean_Widget_MsgEmbed_instRpcEncodableMsgEmbed_enc____x40_Lean_Widget_InteractiveDiagnostic___hyg_551____spec__4(size_t, size_t, lean_object*, lean_object*); static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_InteractiveDiagnostic___hyg_2975____closed__5; @@ -395,7 +392,6 @@ static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget static lean_object* l_Lean_Widget_Lean_Widget_StrictOrLazy_instRpcEncodableStrictOrLazy_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_4____at_Lean_Widget_Lean_Widget_MsgEmbed_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_551____spec__5___closed__2; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_Lean_Lsp_DiagnosticWith_instRpcEncodableDiagnosticWith_enc____x40_Lean_Widget_InteractiveDiagnostic___hyg_1549____spec__2___closed__2; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); -lean_object* lean_dbg_trace(lean_object*, lean_object*); static lean_object* l_Lean_Widget_Lean_Lsp_DiagnosticWith_instRpcEncodableDiagnosticWith_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_1549____rarg___closed__1; lean_object* l_Lean_MessageData_isDeprecationWarning___lambda__1___boxed(lean_object*); static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_Lean_Lsp_DiagnosticWith_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1634____closed__59; @@ -17510,25 +17506,6 @@ lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l_dbgTraceVal___at___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_chopUpChildren___spec__1___lambda__1(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_inc(x_1); -return x_1; -} -} -LEAN_EXPORT lean_object* l_dbgTraceVal___at___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_chopUpChildren___spec__1(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; lean_object* x_4; -lean_inc(x_1); -x_2 = l_Nat_repr(x_1); -x_3 = lean_alloc_closure((void*)(l_dbgTraceVal___at___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_chopUpChildren___spec__1___lambda__1___boxed), 2, 1); -lean_closure_set(x_3, 0, x_1); -x_4 = lean_dbg_trace(x_2, x_3); -return x_4; -} -} static lean_object* _init_l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_chopUpChildren___closed__1() { _start: { @@ -17577,7 +17554,7 @@ return x_8; } else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; 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; lean_object* x_27; lean_object* x_28; +lean_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_26; lean_object* x_27; lean_inc(x_3); x_9 = l_Array_ofSubarray___rarg(x_3); x_10 = l_Subarray_size___rarg(x_3); @@ -17594,39 +17571,28 @@ x_16 = l_Array_ofSubarray___rarg(x_15); x_17 = lean_nat_sub(x_6, x_2); lean_dec(x_2); lean_dec(x_6); -x_18 = l_dbgTraceVal___at___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_chopUpChildren___spec__1(x_17); -x_19 = l_Nat_repr(x_18); -x_20 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_20, 0, x_19); -x_21 = l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_chopUpChildren___closed__1; -x_22 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_20); -x_23 = l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_chopUpChildren___closed__3; -x_24 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_24, 0, x_22); -lean_ctor_set(x_24, 1, x_23); -x_25 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_25, 0, x_24); -x_26 = 1; -x_27 = lean_alloc_ctor(9, 3, 1); -lean_ctor_set(x_27, 0, x_1); -lean_ctor_set(x_27, 1, x_25); -lean_ctor_set(x_27, 2, x_12); -lean_ctor_set_uint8(x_27, sizeof(void*)*3, x_26); -x_28 = lean_array_push(x_16, x_27); -return x_28; -} -} +x_18 = l_Nat_repr(x_17); +x_19 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_19, 0, x_18); +x_20 = l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_chopUpChildren___closed__1; +x_21 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_19); +x_22 = l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_chopUpChildren___closed__3; +x_23 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +x_24 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_24, 0, x_23); +x_25 = 1; +x_26 = lean_alloc_ctor(9, 3, 1); +lean_ctor_set(x_26, 0, x_1); +lean_ctor_set(x_26, 1, x_24); +lean_ctor_set(x_26, 2, x_12); +lean_ctor_set_uint8(x_26, sizeof(void*)*3, x_25); +x_27 = lean_array_push(x_16, x_26); +return x_27; } -LEAN_EXPORT lean_object* l_dbgTraceVal___at___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_chopUpChildren___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l_dbgTraceVal___at___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_chopUpChildren___spec__1___lambda__1(x_1, x_2); -lean_dec(x_2); -lean_dec(x_1); -return x_3; } } LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_go___spec__1(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { From c43a6b53413f1f597eb5144692730a9132c3e473 Mon Sep 17 00:00:00 2001 From: Joe Hendrix Date: Mon, 11 Mar 2024 14:40:48 -0700 Subject: [PATCH 14/32] chore: upstream Std.Data.Int (#3635) This depends on #3634. --- src/Init/Data/BitVec/Basic.lean | 3 + src/Init/Data/BitVec/Lemmas.lean | 8 +- src/Init/Data/Int.lean | 1 + src/Init/Data/Int/DivMod.lean | 6 + src/Init/Data/Int/DivModLemmas.lean | 1011 ++++++++++++++++++--- src/Init/Data/Int/Gcd.lean | 38 + src/Init/Data/Int/Lemmas.lean | 33 - src/Init/Data/Int/Order.lean | 522 +++++++++++ src/Init/Data/Int/Pow.lean | 44 + src/Init/Data/Nat/Basic.lean | 3 + src/Init/Data/Nat/Bitwise/Lemmas.lean | 1 + src/Init/Data/Nat/Div.lean | 7 + src/Init/Data/Nat/Dvd.lean | 9 +- src/Init/Data/Nat/Lemmas.lean | 5 - src/Init/Data/Random.lean | 1 - src/Init/Omega/Int.lean | 1 - src/Init/Omega/IntList.lean | 2 +- src/Lean/Elab/Tactic/Omega/MinNatAbs.lean | 4 +- src/Lean/Elab/Tactic/Omega/OmegaM.lean | 4 +- 19 files changed, 1539 insertions(+), 164 deletions(-) create mode 100644 src/Init/Data/Int/Pow.lean diff --git a/src/Init/Data/BitVec/Basic.lean b/src/Init/Data/BitVec/Basic.lean index 5d48d1843bdf..6558a4d7f525 100644 --- a/src/Init/Data/BitVec/Basic.lean +++ b/src/Init/Data/BitVec/Basic.lean @@ -73,6 +73,9 @@ protected def toNat (a : BitVec n) : Nat := a.toFin.val /-- Return the bound in terms of toNat. -/ theorem isLt (x : BitVec w) : x.toNat < 2^w := x.toFin.isLt +@[deprecated isLt] +theorem toNat_lt (x : BitVec n) : x.toNat < 2^n := x.isLt + /-- Theorem for normalizing the bit vector literal representation. -/ -- TODO: This needs more usage data to assess which direction the simp should go. @[simp, bv_toNat] theorem ofNat_eq_ofNat : @OfNat.ofNat (BitVec n) i _ = .ofNat n i := rfl diff --git a/src/Init/Data/BitVec/Lemmas.lean b/src/Init/Data/BitVec/Lemmas.lean index 72c64bef4d26..7ad648eb102f 100644 --- a/src/Init/Data/BitVec/Lemmas.lean +++ b/src/Init/Data/BitVec/Lemmas.lean @@ -29,8 +29,6 @@ theorem eq_of_toNat_eq {n} : ∀ {i j : BitVec n}, i.toNat = j.toNat → i = j @[bv_toNat] theorem toNat_ne (x y : BitVec n) : x ≠ y ↔ x.toNat ≠ y.toNat := by rw [Ne, toNat_eq] -theorem toNat_lt (x : BitVec n) : x.toNat < 2^n := x.toFin.2 - theorem testBit_toNat (x : BitVec w) : x.toNat.testBit i = x.getLsb i := rfl @[simp] theorem getLsb_ofFin (x : Fin (2^n)) (i : Nat) : @@ -458,12 +456,12 @@ theorem not_def {x : BitVec v} : ~~~x = allOnes v ^^^ x := rfl | y+1 => rw [Nat.succ_eq_add_one] at h rw [← h] - rw [Nat.testBit_two_pow_sub_succ (toNat_lt _)] + rw [Nat.testBit_two_pow_sub_succ (isLt _)] · cases w : decide (i < v) · simp at w simp [w] rw [Nat.testBit_lt_two_pow] - calc BitVec.toNat x < 2 ^ v := toNat_lt _ + calc BitVec.toNat x < 2 ^ v := isLt _ _ ≤ 2 ^ i := Nat.pow_le_pow_of_le_right Nat.zero_lt_two w · simp @@ -520,7 +518,7 @@ theorem shiftLeftZeroExtend_eq {x : BitVec w} : · simp rw [Nat.mod_eq_of_lt] rw [Nat.shiftLeft_eq, Nat.pow_add] - exact Nat.mul_lt_mul_of_pos_right (BitVec.toNat_lt x) (Nat.two_pow_pos _) + exact Nat.mul_lt_mul_of_pos_right x.isLt (Nat.two_pow_pos _) · omega @[simp] theorem getLsb_shiftLeftZeroExtend (x : BitVec m) (n : Nat) : diff --git a/src/Init/Data/Int.lean b/src/Init/Data/Int.lean index def6f7557781..83cec56d5e42 100644 --- a/src/Init/Data/Int.lean +++ b/src/Init/Data/Int.lean @@ -11,3 +11,4 @@ import Init.Data.Int.DivModLemmas import Init.Data.Int.Gcd import Init.Data.Int.Lemmas import Init.Data.Int.Order +import Init.Data.Int.Pow diff --git a/src/Init/Data/Int/DivMod.lean b/src/Init/Data/Int/DivMod.lean index bb4e98adc033..3f9b243d3479 100644 --- a/src/Init/Data/Int/DivMod.lean +++ b/src/Init/Data/Int/DivMod.lean @@ -160,6 +160,12 @@ instance : Mod Int where @[simp, norm_cast] theorem ofNat_ediv (m n : Nat) : (↑(m / n) : Int) = ↑m / ↑n := rfl +theorem ofNat_div (m n : Nat) : ↑(m / n) = div ↑m ↑n := rfl + +theorem ofNat_fdiv : ∀ m n : Nat, ↑(m / n) = fdiv ↑m ↑n + | 0, _ => by simp [fdiv] + | succ _, _ => rfl + /-! # `bmod` ("balanced" mod) diff --git a/src/Init/Data/Int/DivModLemmas.lean b/src/Init/Data/Int/DivModLemmas.lean index f3fb3406abbd..6ef222542037 100644 --- a/src/Init/Data/Int/DivModLemmas.lean +++ b/src/Init/Data/Int/DivModLemmas.lean @@ -8,7 +8,6 @@ prelude import Init.Data.Int.DivMod import Init.Data.Int.Order import Init.Data.Nat.Dvd -import Init.RCases /-! # Lemmas about integer division needed to bootstrap `omega`. @@ -19,7 +18,116 @@ open Nat (succ) namespace Int -/-! ### `/` -/ +/-! ### dvd -/ + +protected theorem dvd_def (a b : Int) : (a ∣ b) = Exists (fun c => b = a * c) := rfl + +protected theorem dvd_zero (n : Int) : n ∣ 0 := ⟨0, (Int.mul_zero _).symm⟩ + +protected theorem dvd_refl (n : Int) : n ∣ n := ⟨1, (Int.mul_one _).symm⟩ + +protected theorem one_dvd (n : Int) : 1 ∣ n := ⟨n, (Int.one_mul n).symm⟩ + +protected theorem dvd_trans : ∀ {a b c : Int}, a ∣ b → b ∣ c → a ∣ c + | _, _, _, ⟨d, rfl⟩, ⟨e, rfl⟩ => Exists.intro (d * e) (by rw [Int.mul_assoc]) + +@[norm_cast] theorem ofNat_dvd {m n : Nat} : (↑m : Int) ∣ ↑n ↔ m ∣ n := by + refine ⟨fun ⟨a, ae⟩ => ?_, fun ⟨k, e⟩ => ⟨k, by rw [e, Int.ofNat_mul]⟩⟩ + match Int.le_total a 0 with + | .inl h => + have := ae.symm ▸ Int.mul_nonpos_of_nonneg_of_nonpos (ofNat_zero_le _) h + rw [Nat.le_antisymm (ofNat_le.1 this) (Nat.zero_le _)] + apply Nat.dvd_zero + | .inr h => match a, eq_ofNat_of_zero_le h with + | _, ⟨k, rfl⟩ => exact ⟨k, Int.ofNat.inj ae⟩ + +theorem dvd_antisymm {a b : Int} (H1 : 0 ≤ a) (H2 : 0 ≤ b) : a ∣ b → b ∣ a → a = b := by + rw [← natAbs_of_nonneg H1, ← natAbs_of_nonneg H2] + rw [ofNat_dvd, ofNat_dvd, ofNat_inj] + apply Nat.dvd_antisymm + +@[simp] protected theorem zero_dvd {n : Int} : 0 ∣ n ↔ n = 0 := + Iff.intro (fun ⟨k, e⟩ => by rw [e, Int.zero_mul]) + (fun h => h.symm ▸ Int.dvd_refl _) + +protected theorem dvd_mul_right (a b : Int) : a ∣ a * b := ⟨_, rfl⟩ + +protected theorem dvd_mul_left (a b : Int) : b ∣ a * b := ⟨_, Int.mul_comm ..⟩ + +protected theorem neg_dvd {a b : Int} : -a ∣ b ↔ a ∣ b := by + constructor <;> exact fun ⟨k, e⟩ => + ⟨-k, by simp [e, Int.neg_mul, Int.mul_neg, Int.neg_neg]⟩ + +protected theorem dvd_neg {a b : Int} : a ∣ -b ↔ a ∣ b := by + constructor <;> exact fun ⟨k, e⟩ => + ⟨-k, by simp [← e, Int.neg_mul, Int.mul_neg, Int.neg_neg]⟩ + +@[simp] theorem natAbs_dvd_natAbs {a b : Int} : natAbs a ∣ natAbs b ↔ a ∣ b := by + refine ⟨fun ⟨k, hk⟩ => ?_, fun ⟨k, hk⟩ => ⟨natAbs k, hk.symm ▸ natAbs_mul a k⟩⟩ + rw [← natAbs_ofNat k, ← natAbs_mul, natAbs_eq_natAbs_iff] at hk + cases hk <;> subst b + · apply Int.dvd_mul_right + · rw [← Int.mul_neg]; apply Int.dvd_mul_right + +theorem ofNat_dvd_left {n : Nat} {z : Int} : (↑n : Int) ∣ z ↔ n ∣ z.natAbs := by + rw [← natAbs_dvd_natAbs, natAbs_ofNat] + +protected theorem dvd_add : ∀ {a b c : Int}, a ∣ b → a ∣ c → a ∣ b + c + | _, _, _, ⟨d, rfl⟩, ⟨e, rfl⟩ => ⟨d + e, by rw [Int.mul_add]⟩ + +protected theorem dvd_sub : ∀ {a b c : Int}, a ∣ b → a ∣ c → a ∣ b - c + | _, _, _, ⟨d, rfl⟩, ⟨e, rfl⟩ => ⟨d - e, by rw [Int.mul_sub]⟩ + +protected theorem dvd_add_left {a b c : Int} (H : a ∣ c) : a ∣ b + c ↔ a ∣ b := + ⟨fun h => by have := Int.dvd_sub h H; rwa [Int.add_sub_cancel] at this, (Int.dvd_add · H)⟩ + +protected theorem dvd_add_right {a b c : Int} (H : a ∣ b) : a ∣ b + c ↔ a ∣ c := by + rw [Int.add_comm, Int.dvd_add_left H] + +protected theorem dvd_iff_dvd_of_dvd_sub {a b c : Int} (H : a ∣ b - c) : a ∣ b ↔ a ∣ c := + ⟨fun h => Int.sub_sub_self b c ▸ Int.dvd_sub h H, + fun h => Int.sub_add_cancel b c ▸ Int.dvd_add H h⟩ + +protected theorem dvd_iff_dvd_of_dvd_add {a b c : Int} (H : a ∣ b + c) : a ∣ b ↔ a ∣ c := by + rw [← Int.sub_neg] at H; rw [Int.dvd_iff_dvd_of_dvd_sub H, Int.dvd_neg] + +theorem le_of_dvd {a b : Int} (bpos : 0 < b) (H : a ∣ b) : a ≤ b := + match a, b, eq_succ_of_zero_lt bpos, H with + | ofNat _, _, ⟨n, rfl⟩, H => ofNat_le.2 <| Nat.le_of_dvd n.succ_pos <| ofNat_dvd.1 H + | -[_+1], _, ⟨_, rfl⟩, _ => Int.le_trans (Int.le_of_lt <| negSucc_lt_zero _) (ofNat_zero_le _) + +theorem natAbs_dvd {a b : Int} : (a.natAbs : Int) ∣ b ↔ a ∣ b := + match natAbs_eq a with + | .inl e => by rw [← e] + | .inr e => by rw [← Int.neg_dvd, ← e] + +theorem dvd_natAbs {a b : Int} : a ∣ b.natAbs ↔ a ∣ b := + match natAbs_eq b with + | .inl e => by rw [← e] + | .inr e => by rw [← Int.dvd_neg, ← e] + +theorem natAbs_dvd_self {a : Int} : (a.natAbs : Int) ∣ a := by + rw [Int.natAbs_dvd] + exact Int.dvd_refl a + +theorem dvd_natAbs_self {a : Int} : a ∣ (a.natAbs : Int) := by + rw [Int.dvd_natAbs] + exact Int.dvd_refl a + +theorem ofNat_dvd_right {n : Nat} {z : Int} : z ∣ (↑n : Int) ↔ z.natAbs ∣ n := by + rw [← natAbs_dvd_natAbs, natAbs_ofNat] + +theorem eq_one_of_dvd_one {a : Int} (H : 0 ≤ a) (H' : a ∣ 1) : a = 1 := + match a, eq_ofNat_of_zero_le H, H' with + | _, ⟨_, rfl⟩, H' => congrArg ofNat <| Nat.eq_one_of_dvd_one <| ofNat_dvd.1 H' + +theorem eq_one_of_mul_eq_one_right {a b : Int} (H : 0 ≤ a) (H' : a * b = 1) : a = 1 := + eq_one_of_dvd_one H ⟨b, H'.symm⟩ + +theorem eq_one_of_mul_eq_one_left {a b : Int} (H : 0 ≤ b) (H' : a * b = 1) : b = 1 := + eq_one_of_mul_eq_one_right H <| by rw [Int.mul_comm, H'] + +/-! ### *div zero -/ @[simp] theorem zero_ediv : ∀ b : Int, 0 / b = 0 | ofNat _ => show ofNat _ = _ by simp @@ -29,14 +137,173 @@ namespace Int | ofNat _ => show ofNat _ = _ by simp | -[_+1] => rfl +@[simp] protected theorem zero_div : ∀ b : Int, div 0 b = 0 + | ofNat _ => show ofNat _ = _ by simp + | -[_+1] => show -ofNat _ = _ by simp + +@[simp] protected theorem div_zero : ∀ a : Int, div a 0 = 0 + | ofNat _ => show ofNat _ = _ by simp + | -[_+1] => rfl + +@[simp] theorem zero_fdiv (b : Int) : fdiv 0 b = 0 := by cases b <;> rfl + +@[simp] protected theorem fdiv_zero : ∀ a : Int, fdiv a 0 = 0 + | 0 => rfl + | succ _ => rfl + | -[_+1] => rfl + +/-! ### div equivalences -/ + +theorem div_eq_ediv : ∀ {a b : Int}, 0 ≤ a → 0 ≤ b → a.div b = a / b + | 0, _, _, _ | _, 0, _, _ => by simp + | succ _, succ _, _, _ => rfl + +theorem fdiv_eq_ediv : ∀ (a : Int) {b : Int}, 0 ≤ b → fdiv a b = a / b + | 0, _, _ | -[_+1], 0, _ => by simp + | succ _, ofNat _, _ | -[_+1], succ _, _ => rfl + +theorem fdiv_eq_div {a b : Int} (Ha : 0 ≤ a) (Hb : 0 ≤ b) : fdiv a b = div a b := + div_eq_ediv Ha Hb ▸ fdiv_eq_ediv _ Hb + +/-! ### mod zero -/ + +@[simp] theorem zero_emod (b : Int) : 0 % b = 0 := rfl + +@[simp] theorem emod_zero : ∀ a : Int, a % 0 = a + | ofNat _ => congrArg ofNat <| Nat.mod_zero _ + | -[_+1] => congrArg negSucc <| Nat.mod_zero _ + +@[simp] theorem zero_mod (b : Int) : mod 0 b = 0 := by cases b <;> simp [mod] + +@[simp] theorem mod_zero : ∀ a : Int, mod a 0 = a + | ofNat _ => congrArg ofNat <| Nat.mod_zero _ + | -[_+1] => rfl + +@[simp] theorem zero_fmod (b : Int) : fmod 0 b = 0 := by cases b <;> rfl + +@[simp] theorem fmod_zero : ∀ a : Int, fmod a 0 = a + | 0 => rfl + | succ _ => congrArg ofNat <| Nat.mod_zero _ + | -[_+1] => congrArg negSucc <| Nat.mod_zero _ + +/-! ### ofNat mod -/ + +@[simp, norm_cast] theorem ofNat_emod (m n : Nat) : (↑(m % n) : Int) = m % n := rfl + + +/-! ### mod definitiions -/ + +theorem emod_add_ediv : ∀ a b : Int, a % b + b * (a / b) = a + | ofNat _, ofNat _ => congrArg ofNat <| Nat.mod_add_div .. + | ofNat m, -[n+1] => by + show (m % succ n + -↑(succ n) * -↑(m / succ n) : Int) = m + rw [Int.neg_mul_neg]; exact congrArg ofNat <| Nat.mod_add_div .. + | -[_+1], 0 => by rw [emod_zero]; rfl + | -[m+1], succ n => aux m n.succ + | -[m+1], -[n+1] => aux m n.succ +where + aux (m n : Nat) : n - (m % n + 1) - (n * (m / n) + n) = -[m+1] := by + rw [← ofNat_emod, ← ofNat_ediv, ← Int.sub_sub, negSucc_eq, Int.sub_sub n, + ← Int.neg_neg (_-_), Int.neg_sub, Int.sub_sub_self, Int.add_right_comm] + exact congrArg (fun x => -(ofNat x + 1)) (Nat.mod_add_div ..) + +theorem emod_add_ediv' (a b : Int) : a % b + a / b * b = a := by + rw [Int.mul_comm]; exact emod_add_ediv .. + +theorem ediv_add_emod (a b : Int) : b * (a / b) + a % b = a := by + rw [Int.add_comm]; exact emod_add_ediv .. + +theorem ediv_add_emod' (a b : Int) : a / b * b + a % b = a := by + rw [Int.mul_comm]; exact ediv_add_emod .. + +theorem emod_def (a b : Int) : a % b = a - b * (a / b) := by + rw [← Int.add_sub_cancel (a % b), emod_add_ediv] + +theorem mod_add_div : ∀ a b : Int, mod a b + b * (a.div b) = a + | ofNat _, ofNat _ => congrArg ofNat (Nat.mod_add_div ..) + | ofNat m, -[n+1] => by + show (m % succ n + -↑(succ n) * -↑(m / succ n) : Int) = m + rw [Int.neg_mul_neg]; exact congrArg ofNat (Nat.mod_add_div ..) + | -[_+1], 0 => rfl + | -[m+1], ofNat n => by + show -(↑((succ m) % n) : Int) + ↑n * -↑(succ m / n) = -↑(succ m) + rw [Int.mul_neg, ← Int.neg_add] + exact congrArg (-ofNat ·) (Nat.mod_add_div ..) + | -[m+1], -[n+1] => by + show -(↑(succ m % succ n) : Int) + -↑(succ n) * ↑(succ m / succ n) = -↑(succ m) + rw [Int.neg_mul, ← Int.neg_add] + exact congrArg (-ofNat ·) (Nat.mod_add_div ..) + +theorem div_add_mod (a b : Int) : b * a.div b + mod a b = a := by + rw [Int.add_comm]; apply mod_add_div .. + +theorem mod_add_div' (m k : Int) : mod m k + m.div k * k = m := by + rw [Int.mul_comm]; apply mod_add_div + +theorem div_add_mod' (m k : Int) : m.div k * k + mod m k = m := by + rw [Int.mul_comm]; apply div_add_mod + +theorem mod_def (a b : Int) : mod a b = a - b * a.div b := by + rw [← Int.add_sub_cancel (mod a b), mod_add_div] + +theorem fmod_add_fdiv : ∀ a b : Int, a.fmod b + b * a.fdiv b = a + | 0, ofNat _ | 0, -[_+1] => congrArg ofNat <| by simp + | succ m, ofNat n => congrArg ofNat <| Nat.mod_add_div .. + | succ m, -[n+1] => by + show subNatNat (m % succ n) n + (↑(succ n * (m / succ n)) + n + 1) = (m + 1) + rw [Int.add_comm _ n, ← Int.add_assoc, ← Int.add_assoc, + Int.subNatNat_eq_coe, Int.sub_add_cancel] + exact congrArg (ofNat · + 1) <| Nat.mod_add_div .. + | -[_+1], 0 => by rw [fmod_zero]; rfl + | -[m+1], succ n => by + show subNatNat .. - (↑(succ n * (m / succ n)) + ↑(succ n)) = -↑(succ m) + rw [Int.subNatNat_eq_coe, ← Int.sub_sub, ← Int.neg_sub, Int.sub_sub, Int.sub_sub_self] + exact congrArg (-ofNat ·) <| Nat.succ_add .. ▸ Nat.mod_add_div .. ▸ rfl + | -[m+1], -[n+1] => by + show -(↑(succ m % succ n) : Int) + -↑(succ n * (succ m / succ n)) = -↑(succ m) + rw [← Int.neg_add]; exact congrArg (-ofNat ·) <| Nat.mod_add_div .. + +theorem fdiv_add_fmod (a b : Int) : b * a.fdiv b + a.fmod b = a := by + rw [Int.add_comm]; exact fmod_add_fdiv .. + +theorem fmod_def (a b : Int) : a.fmod b = a - b * a.fdiv b := by + rw [← Int.add_sub_cancel (a.fmod b), fmod_add_fdiv] + +/-! ### mod equivalences -/ + +theorem fmod_eq_emod (a : Int) {b : Int} (hb : 0 ≤ b) : fmod a b = a % b := by + simp [fmod_def, emod_def, fdiv_eq_ediv _ hb] + +theorem mod_eq_emod {a b : Int} (ha : 0 ≤ a) (hb : 0 ≤ b) : mod a b = a % b := by + simp [emod_def, mod_def, div_eq_ediv ha hb] + +theorem fmod_eq_mod {a b : Int} (Ha : 0 ≤ a) (Hb : 0 ≤ b) : fmod a b = mod a b := + mod_eq_emod Ha Hb ▸ fmod_eq_emod _ Hb + +/-! ### `/` ediv -/ @[simp] protected theorem ediv_neg : ∀ a b : Int, a / (-b) = -(a / b) | ofNat m, 0 => show ofNat (m / 0) = -↑(m / 0) by rw [Nat.div_zero]; rfl | ofNat m, -[n+1] => (Int.neg_neg _).symm | ofNat m, succ n | -[m+1], 0 | -[m+1], succ n | -[m+1], -[n+1] => rfl +theorem ediv_neg' {a b : Int} (Ha : a < 0) (Hb : 0 < b) : a / b < 0 := + match a, b, eq_negSucc_of_lt_zero Ha, eq_succ_of_zero_lt Hb with + | _, _, ⟨_, rfl⟩, ⟨_, rfl⟩ => negSucc_lt_zero _ + protected theorem div_def (a b : Int) : a / b = Int.ediv a b := rfl +theorem negSucc_ediv (m : Nat) {b : Int} (H : 0 < b) : -[m+1] / b = -(div m b + 1) := + match b, eq_succ_of_zero_lt H with + | _, ⟨_, rfl⟩ => rfl + +theorem ediv_nonneg {a b : Int} (Ha : 0 ≤ a) (Hb : 0 ≤ b) : 0 ≤ a / b := + match a, b, eq_ofNat_of_zero_le Ha, eq_ofNat_of_zero_le Hb with + | _, _, ⟨_, rfl⟩, ⟨_, rfl⟩ => ofNat_zero_le _ + +theorem ediv_nonpos {a b : Int} (Ha : 0 ≤ a) (Hb : b ≤ 0) : a / b ≤ 0 := + Int.nonpos_of_neg_nonneg <| Int.ediv_neg .. ▸ Int.ediv_nonneg Ha (Int.neg_nonneg_of_nonpos Hb) + theorem add_mul_ediv_right (a b : Int) {c : Int} (H : c ≠ 0) : (a + b * c) / c = a / c + b := suffices ∀ {{a b c : Int}}, 0 < c → (a + b * c).ediv c = a.ediv c + b from match Int.lt_trichotomy c 0 with @@ -54,12 +321,11 @@ theorem add_mul_ediv_right (a b : Int) {c : Int} (H : c ≠ 0) : (a + b * c) / c | ofNat m => congrArg ofNat <| Nat.add_mul_div_right _ _ k.succ_pos | -[m+1] => by show ((n * k.succ : Nat) - m.succ : Int).ediv k.succ = n - (m / k.succ + 1 : Nat) - if h : m < n * k.succ then - rw [← Int.ofNat_sub h, ← Int.ofNat_sub ((Nat.div_lt_iff_lt_mul k.succ_pos).2 h)] + by_cases h : m < n * k.succ + · rw [← Int.ofNat_sub h, ← Int.ofNat_sub ((Nat.div_lt_iff_lt_mul k.succ_pos).2 h)] apply congrArg ofNat rw [Nat.mul_comm, Nat.mul_sub_div]; rwa [Nat.mul_comm] - else - have h := Nat.not_lt.1 h + · have h := Nat.not_lt.1 h have H {a b : Nat} (h : a ≤ b) : (a : Int) + -((b : Int) + 1) = -[b - a +1] := by rw [negSucc_eq, Int.ofNat_sub h] simp only [Int.sub_eq_add_neg, Int.neg_add, Int.neg_neg, Int.add_left_comm, Int.add_assoc] @@ -91,42 +357,64 @@ theorem div_nonneg_iff_of_pos {a b : Int} (h : 0 < b) : a / b ≥ 0 ↔ a ≥ 0 rcases a with ⟨a⟩ <;> simp [Int.ediv] exact decide_eq_decide.mp rfl -/-! ### mod -/ +theorem ediv_eq_zero_of_lt {a b : Int} (H1 : 0 ≤ a) (H2 : a < b) : a / b = 0 := + match a, b, eq_ofNat_of_zero_le H1, eq_succ_of_zero_lt (Int.lt_of_le_of_lt H1 H2) with + | _, _, ⟨_, rfl⟩, ⟨_, rfl⟩ => congrArg Nat.cast <| Nat.div_eq_of_lt <| ofNat_lt.1 H2 + +theorem add_mul_ediv_left (a : Int) {b : Int} + (c : Int) (H : b ≠ 0) : (a + b * c) / b = a / b + c := + Int.mul_comm .. ▸ Int.add_mul_ediv_right _ _ H + +@[simp] theorem mul_ediv_mul_of_pos {a : Int} + (b c : Int) (H : 0 < a) : (a * b) / (a * c) = b / c := + suffices ∀ (m k : Nat) (b : Int), (m.succ * b) / (m.succ * k) = b / k from + match a, eq_succ_of_zero_lt H, c, Int.eq_nat_or_neg c with + | _, ⟨m, rfl⟩, _, ⟨k, .inl rfl⟩ => this _ .. + | _, ⟨m, rfl⟩, _, ⟨k, .inr rfl⟩ => by + rw [Int.mul_neg, Int.ediv_neg, Int.ediv_neg]; apply congrArg Neg.neg; apply this + fun m k b => + match b, k with + | ofNat n, k => congrArg ofNat (Nat.mul_div_mul_left _ _ m.succ_pos) + | -[n+1], 0 => by + rw [Int.ofNat_zero, Int.mul_zero, Int.ediv_zero, Int.ediv_zero] + | -[n+1], succ k => congrArg negSucc <| + show (m.succ * n + m) / (m.succ * k.succ) = n / k.succ by + apply Nat.div_eq_of_lt_le + · refine Nat.le_trans ?_ (Nat.le_add_right _ _) + rw [← Nat.mul_div_mul_left _ _ m.succ_pos] + apply Nat.div_mul_le_self + · show m.succ * n.succ ≤ _ + rw [Nat.mul_left_comm] + apply Nat.mul_le_mul_left + apply (Nat.div_lt_iff_lt_mul k.succ_pos).1 + apply Nat.lt_succ_self + +@[simp] theorem mul_ediv_mul_of_pos_left + (a : Int) {b : Int} (c : Int) (H : 0 < b) : (a * b) / (c * b) = a / c := by + rw [Int.mul_comm, Int.mul_comm c, mul_ediv_mul_of_pos _ _ H] + +protected theorem ediv_eq_of_eq_mul_right {a b c : Int} + (H1 : b ≠ 0) (H2 : a = b * c) : a / b = c := by rw [H2, Int.mul_ediv_cancel_left _ H1] + +protected theorem eq_ediv_of_mul_eq_right {a b c : Int} + (H1 : a ≠ 0) (H2 : a * b = c) : b = c / a := + (Int.ediv_eq_of_eq_mul_right H1 H2.symm).symm + +protected theorem ediv_eq_of_eq_mul_left {a b c : Int} + (H1 : b ≠ 0) (H2 : a = c * b) : a / b = c := + Int.ediv_eq_of_eq_mul_right H1 (by rw [Int.mul_comm, H2]) + +/-! ### emod -/ theorem mod_def' (m n : Int) : m % n = emod m n := rfl -theorem ofNat_mod (m n : Nat) : (↑(m % n) : Int) = mod m n := rfl +theorem negSucc_emod (m : Nat) {b : Int} (bpos : 0 < b) : -[m+1] % b = b - 1 - m % b := by + rw [Int.sub_sub, Int.add_comm] + match b, eq_succ_of_zero_lt bpos with + | _, ⟨n, rfl⟩ => rfl theorem ofNat_mod_ofNat (m n : Nat) : (m % n : Int) = ↑(m % n) := rfl -@[simp, norm_cast] theorem ofNat_emod (m n : Nat) : (↑(m % n) : Int) = m % n := rfl - -@[simp] theorem zero_emod (b : Int) : 0 % b = 0 := by simp [mod_def', emod] - -@[simp] theorem emod_zero : ∀ a : Int, a % 0 = a - | ofNat _ => congrArg ofNat <| Nat.mod_zero _ - | -[_+1] => congrArg negSucc <| Nat.mod_zero _ - -theorem emod_add_ediv : ∀ a b : Int, a % b + b * (a / b) = a - | ofNat _, ofNat _ => congrArg ofNat <| Nat.mod_add_div .. - | ofNat m, -[n+1] => by - show (m % succ n + -↑(succ n) * -↑(m / succ n) : Int) = m - rw [Int.neg_mul_neg]; exact congrArg ofNat <| Nat.mod_add_div .. - | -[_+1], 0 => by rw [emod_zero]; rfl - | -[m+1], succ n => aux m n.succ - | -[m+1], -[n+1] => aux m n.succ -where - aux (m n : Nat) : n - (m % n + 1) - (n * (m / n) + n) = -[m+1] := by - rw [← ofNat_emod, ← ofNat_ediv, ← Int.sub_sub, negSucc_eq, Int.sub_sub n, - ← Int.neg_neg (_-_), Int.neg_sub, Int.sub_sub_self, Int.add_right_comm] - exact congrArg (fun x => -(ofNat x + 1)) (Nat.mod_add_div ..) - -theorem ediv_add_emod (a b : Int) : b * (a / b) + a % b = a := - (Int.add_comm ..).trans (emod_add_ediv ..) - -theorem emod_def (a b : Int) : a % b = a - b * (a / b) := by - rw [← Int.add_sub_cancel (a % b), emod_add_ediv] - theorem emod_nonneg : ∀ (a : Int) {b : Int}, b ≠ 0 → 0 ≤ a % b | ofNat _, _, _ => ofNat_zero_le _ | -[_+1], _, H => Int.sub_nonneg_of_le <| ofNat_le.2 <| Nat.mod_lt _ (natAbs_pos.2 H) @@ -146,9 +434,6 @@ theorem lt_mul_ediv_self_add {x k : Int} (h : 0 < k) : x < k * (x / k) + k := _ = k * (x / k) + x % k := (ediv_add_emod _ _).symm _ < k * (x / k) + k := Int.add_lt_add_left (emod_lt_of_pos x h) _ -theorem emod_add_ediv' (m k : Int) : m % k + m / k * k = m := by - rw [Int.mul_comm]; apply emod_add_ediv - @[simp] theorem add_mul_emod_self {a b c : Int} : (a + b * c) % c = a % c := if cz : c = 0 then by rw [cz, Int.mul_zero, Int.add_zero] @@ -168,6 +453,9 @@ theorem emod_add_ediv' (m k : Int) : m % k + m / k * k = m := by theorem neg_emod {a b : Int} : -a % b = (b - a) % b := by rw [← add_emod_self_left]; rfl +@[simp] theorem emod_neg (a b : Int) : a % -b = a % b := by + rw [emod_def, emod_def, Int.ediv_neg, Int.neg_mul_neg] + @[simp] theorem emod_add_emod (m n k : Int) : (m % n + k) % n = (m + k) % n := by have := (add_mul_emod_self_left (m % n + k) n (m / n)).symm rwa [Int.add_right_comm, emod_add_ediv] at this @@ -216,6 +504,14 @@ theorem sub_emod (a b n : Int) : (a - b) % n = (a % n - b % n) % n := by apply (emod_add_cancel_right b).mp rw [Int.sub_add_cancel, ← Int.add_emod_emod, Int.sub_add_cancel, emod_emod] +theorem emod_eq_of_lt {a b : Int} (H1 : 0 ≤ a) (H2 : a < b) : a % b = a := + have b0 := Int.le_trans H1 (Int.le_of_lt H2) + match a, b, eq_ofNat_of_zero_le H1, eq_ofNat_of_zero_le b0 with + | _, _, ⟨_, rfl⟩, ⟨_, rfl⟩ => congrArg ofNat <| Nat.mod_eq_of_lt (Int.ofNat_lt.1 H2) + +@[simp] theorem emod_self_add_one {x : Int} (h : 0 ≤ x) : x % (x + 1) = x := + emod_eq_of_lt h (Int.lt_succ x) + /-! ### properties of `/` and `%` -/ theorem mul_ediv_cancel_of_emod_eq_zero {a b : Int} (H : a % b = 0) : b * (a / b) = a := by @@ -224,58 +520,58 @@ theorem mul_ediv_cancel_of_emod_eq_zero {a b : Int} (H : a % b = 0) : b * (a / b theorem ediv_mul_cancel_of_emod_eq_zero {a b : Int} (H : a % b = 0) : a / b * b = a := by rw [Int.mul_comm, mul_ediv_cancel_of_emod_eq_zero H] -/-! ### dvd -/ - -protected theorem dvd_zero (n : Int) : n ∣ 0 := ⟨0, (Int.mul_zero _).symm⟩ - -protected theorem dvd_refl (n : Int) : n ∣ n := ⟨1, (Int.mul_one _).symm⟩ - -protected theorem one_dvd (n : Int) : 1 ∣ n := ⟨n, (Int.one_mul n).symm⟩ - -protected theorem dvd_trans : ∀ {a b c : Int}, a ∣ b → b ∣ c → a ∣ c - | _, _, _, ⟨d, rfl⟩, ⟨e, rfl⟩ => ⟨d * e, by rw [Int.mul_assoc]⟩ - -@[simp] protected theorem zero_dvd {n : Int} : 0 ∣ n ↔ n = 0 := - ⟨fun ⟨k, e⟩ => by rw [e, Int.zero_mul], fun h => h.symm ▸ Int.dvd_refl _⟩ - -protected theorem neg_dvd {a b : Int} : -a ∣ b ↔ a ∣ b := by - constructor <;> exact fun ⟨k, e⟩ => - ⟨-k, by simp [e, Int.neg_mul, Int.mul_neg, Int.neg_neg]⟩ - -protected theorem dvd_neg {a b : Int} : a ∣ -b ↔ a ∣ b := by - constructor <;> exact fun ⟨k, e⟩ => - ⟨-k, by simp [← e, Int.neg_mul, Int.mul_neg, Int.neg_neg]⟩ - -protected theorem dvd_mul_right (a b : Int) : a ∣ a * b := ⟨_, rfl⟩ - -protected theorem dvd_mul_left (a b : Int) : b ∣ a * b := ⟨_, Int.mul_comm ..⟩ - -protected theorem dvd_add : ∀ {a b c : Int}, a ∣ b → a ∣ c → a ∣ b + c - | _, _, _, ⟨d, rfl⟩, ⟨e, rfl⟩ => ⟨d + e, by rw [Int.mul_add]⟩ - -protected theorem dvd_sub : ∀ {a b c : Int}, a ∣ b → a ∣ c → a ∣ b - c - | _, _, _, ⟨d, rfl⟩, ⟨e, rfl⟩ => ⟨d - e, by rw [Int.mul_sub]⟩ - - -@[norm_cast] theorem ofNat_dvd {m n : Nat} : (↑m : Int) ∣ ↑n ↔ m ∣ n := by - refine ⟨fun ⟨a, ae⟩ => ?_, fun ⟨k, e⟩ => ⟨k, by rw [e, Int.ofNat_mul]⟩⟩ - match Int.le_total a 0 with - | .inl h => - have := ae.symm ▸ Int.mul_nonpos_of_nonneg_of_nonpos (ofNat_zero_le _) h - rw [Nat.le_antisymm (ofNat_le.1 this) (Nat.zero_le _)] - apply Nat.dvd_zero - | .inr h => match a, eq_ofNat_of_zero_le h with - | _, ⟨k, rfl⟩ => exact ⟨k, Int.ofNat.inj ae⟩ - -@[simp] theorem natAbs_dvd_natAbs {a b : Int} : natAbs a ∣ natAbs b ↔ a ∣ b := by - refine ⟨fun ⟨k, hk⟩ => ?_, fun ⟨k, hk⟩ => ⟨natAbs k, hk.symm ▸ natAbs_mul a k⟩⟩ - rw [← natAbs_ofNat k, ← natAbs_mul, natAbs_eq_natAbs_iff] at hk - cases hk <;> subst b - · apply Int.dvd_mul_right - · rw [← Int.mul_neg]; apply Int.dvd_mul_right +theorem emod_two_eq (x : Int) : x % 2 = 0 ∨ x % 2 = 1 := by + have h₁ : 0 ≤ x % 2 := Int.emod_nonneg x (by decide) + have h₂ : x % 2 < 2 := Int.emod_lt_of_pos x (by decide) + match x % 2, h₁, h₂ with + | 0, _, _ => simp + | 1, _, _ => simp + +theorem add_emod_eq_add_emod_left {m n k : Int} (i : Int) + (H : m % n = k % n) : (i + m) % n = (i + k) % n := by + rw [Int.add_comm, add_emod_eq_add_emod_right _ H, Int.add_comm] + +theorem emod_add_cancel_left {m n k i : Int} : (i + m) % n = (i + k) % n ↔ m % n = k % n := by + rw [Int.add_comm, Int.add_comm i, emod_add_cancel_right] + +theorem emod_sub_cancel_right {m n k : Int} (i) : (m - i) % n = (k - i) % n ↔ m % n = k % n := + emod_add_cancel_right _ + +theorem emod_eq_emod_iff_emod_sub_eq_zero {m n k : Int} : m % n = k % n ↔ (m - k) % n = 0 := + (emod_sub_cancel_right k).symm.trans <| by simp [Int.sub_self] + +protected theorem ediv_emod_unique {a b r q : Int} (h : 0 < b) : + a / b = q ∧ a % b = r ↔ r + b * q = a ∧ 0 ≤ r ∧ r < b := by + constructor + · intro ⟨rfl, rfl⟩ + exact ⟨emod_add_ediv a b, emod_nonneg _ (Int.ne_of_gt h), emod_lt_of_pos _ h⟩ + · intro ⟨rfl, hz, hb⟩ + constructor + · rw [Int.add_mul_ediv_left r q (Int.ne_of_gt h), ediv_eq_zero_of_lt hz hb] + simp [Int.zero_add] + · rw [add_mul_emod_self_left, emod_eq_of_lt hz hb] + +@[simp] theorem mul_emod_mul_of_pos + {a : Int} (b c : Int) (H : 0 < a) : (a * b) % (a * c) = a * (b % c) := by + rw [emod_def, emod_def, mul_ediv_mul_of_pos _ _ H, Int.mul_sub, Int.mul_assoc] + +theorem lt_ediv_add_one_mul_self (a : Int) {b : Int} (H : 0 < b) : a < (a / b + 1) * b := by + rw [Int.add_mul, Int.one_mul, Int.mul_comm] + exact Int.lt_add_of_sub_left_lt <| Int.emod_def .. ▸ emod_lt_of_pos _ H + +theorem natAbs_div_le_natAbs (a b : Int) : natAbs (a / b) ≤ natAbs a := + match b, eq_nat_or_neg b with + | _, ⟨n, .inl rfl⟩ => aux _ _ + | _, ⟨n, .inr rfl⟩ => by rw [Int.ediv_neg, natAbs_neg]; apply aux +where + aux : ∀ (a : Int) (n : Nat), natAbs (a / n) ≤ natAbs a + | ofNat _, _ => Nat.div_le_self .. + | -[_+1], 0 => Nat.zero_le _ + | -[_+1], succ _ => Nat.succ_le_succ (Nat.div_le_self _ _) -theorem ofNat_dvd_left {n : Nat} {z : Int} : (↑n : Int) ∣ z ↔ n ∣ z.natAbs := by - rw [← natAbs_dvd_natAbs, natAbs_ofNat] +theorem ediv_le_self {a : Int} (b : Int) (Ha : 0 ≤ a) : a / b ≤ a := by + have := Int.le_trans le_natAbs (ofNat_le.2 <| natAbs_div_le_natAbs a b) + rwa [natAbs_of_nonneg Ha] at this theorem dvd_of_emod_eq_zero {a b : Int} (H : b % a = 0) : a ∣ b := ⟨b / a, (mul_ediv_cancel_of_emod_eq_zero H).symm⟩ @@ -290,19 +586,14 @@ theorem emod_eq_zero_of_dvd : ∀ {a b : Int}, a ∣ b → b % a = 0 theorem dvd_iff_emod_eq_zero (a b : Int) : a ∣ b ↔ b % a = 0 := ⟨emod_eq_zero_of_dvd, dvd_of_emod_eq_zero⟩ -theorem emod_pos_of_not_dvd {a b : Int} (h : ¬ a ∣ b) : a = 0 ∨ 0 < b % a := by - rw [dvd_iff_emod_eq_zero] at h - if w : a = 0 then simp_all - else exact Or.inr (Int.lt_iff_le_and_ne.mpr ⟨emod_nonneg b w, Ne.symm h⟩) - instance decidableDvd : DecidableRel (α := Int) (· ∣ ·) := fun _ _ => decidable_of_decidable_of_iff (dvd_iff_emod_eq_zero ..).symm -protected theorem ediv_mul_cancel {a b : Int} (H : b ∣ a) : a / b * b = a := - ediv_mul_cancel_of_emod_eq_zero (emod_eq_zero_of_dvd H) - -protected theorem mul_ediv_cancel' {a b : Int} (H : a ∣ b) : a * (b / a) = b := by - rw [Int.mul_comm, Int.ediv_mul_cancel H] +theorem emod_pos_of_not_dvd {a b : Int} (h : ¬ a ∣ b) : a = 0 ∨ 0 < b % a := by + rw [dvd_iff_emod_eq_zero] at h + by_cases w : a = 0 + · simp_all + · exact Or.inr (Int.lt_iff_le_and_ne.mpr ⟨emod_nonneg b w, Ne.symm h⟩) protected theorem mul_ediv_assoc (a : Int) : ∀ {b c : Int}, c ∣ b → (a * b) / c = a * (b / c) | _, c, ⟨d, rfl⟩ => @@ -314,8 +605,10 @@ protected theorem mul_ediv_assoc' (b : Int) {a c : Int} rw [Int.mul_comm, Int.mul_ediv_assoc _ h, Int.mul_comm] theorem neg_ediv_of_dvd : ∀ {a b : Int}, b ∣ a → (-a) / b = -(a / b) - | _, b, ⟨c, rfl⟩ => by if bz : b = 0 then simp [bz] else - rw [Int.neg_mul_eq_mul_neg, Int.mul_ediv_cancel_left _ bz, Int.mul_ediv_cancel_left _ bz] + | _, b, ⟨c, rfl⟩ => by + by_cases bz : b = 0 + · simp [bz] + · rw [Int.neg_mul_eq_mul_neg, Int.mul_ediv_cancel_left _ bz, Int.mul_ediv_cancel_left _ bz] theorem sub_ediv_of_dvd (a : Int) {b c : Int} (hcb : c ∣ b) : (a - b) / c = a / c - b / c := by @@ -334,13 +627,389 @@ theorem sub_ediv_of_dvd (a : Int) {b c : Int} @[simp] theorem Int.emod_sub_cancel (x y : Int): (x - y)%y = x%y := by - if h : y = 0 then - simp [h] - else - simp only [Int.emod_def, Int.sub_ediv_of_dvd, Int.dvd_refl, Int.ediv_self h, Int.mul_sub] + by_cases h : y = 0 + · simp [h] + · simp only [Int.emod_def, Int.sub_ediv_of_dvd, Int.dvd_refl, Int.ediv_self h, Int.mul_sub] simp [Int.mul_one, Int.sub_sub, Int.add_comm y] -/-! bmod -/ +/-- If `a % b = c` then `b` divides `a - c`. -/ +theorem dvd_sub_of_emod_eq {a b c : Int} (h : a % b = c) : b ∣ a - c := by + have hx : (a % b) % b = c % b := by + rw [h] + rw [Int.emod_emod, ← emod_sub_cancel_right c, Int.sub_self, zero_emod] at hx + exact dvd_of_emod_eq_zero hx + +protected theorem ediv_mul_cancel {a b : Int} (H : b ∣ a) : a / b * b = a := + ediv_mul_cancel_of_emod_eq_zero (emod_eq_zero_of_dvd H) + +protected theorem mul_ediv_cancel' {a b : Int} (H : a ∣ b) : a * (b / a) = b := by + rw [Int.mul_comm, Int.ediv_mul_cancel H] + +protected theorem eq_mul_of_ediv_eq_right {a b c : Int} + (H1 : b ∣ a) (H2 : a / b = c) : a = b * c := by rw [← H2, Int.mul_ediv_cancel' H1] + +protected theorem ediv_eq_iff_eq_mul_right {a b c : Int} + (H : b ≠ 0) (H' : b ∣ a) : a / b = c ↔ a = b * c := + ⟨Int.eq_mul_of_ediv_eq_right H', Int.ediv_eq_of_eq_mul_right H⟩ + +protected theorem ediv_eq_iff_eq_mul_left {a b c : Int} + (H : b ≠ 0) (H' : b ∣ a) : a / b = c ↔ a = c * b := by + rw [Int.mul_comm]; exact Int.ediv_eq_iff_eq_mul_right H H' + +protected theorem eq_mul_of_ediv_eq_left {a b c : Int} + (H1 : b ∣ a) (H2 : a / b = c) : a = c * b := by + rw [Int.mul_comm, Int.eq_mul_of_ediv_eq_right H1 H2] + +protected theorem eq_zero_of_ediv_eq_zero {d n : Int} (h : d ∣ n) (H : n / d = 0) : n = 0 := by + rw [← Int.mul_ediv_cancel' h, H, Int.mul_zero] + +theorem sub_ediv_of_dvd_sub {a b c : Int} + (hcab : c ∣ a - b) : (a - b) / c = a / c - b / c := by + rw [← Int.add_sub_cancel ((a-b) / c), ← Int.add_ediv_of_dvd_left hcab, Int.sub_add_cancel] + +@[simp] protected theorem ediv_left_inj {a b d : Int} + (hda : d ∣ a) (hdb : d ∣ b) : a / d = b / d ↔ a = b := by + refine ⟨fun h => ?_, congrArg (ediv · d)⟩ + rw [← Int.mul_ediv_cancel' hda, ← Int.mul_ediv_cancel' hdb, h] + +theorem ediv_sign : ∀ a b, a / sign b = a * sign b + | _, succ _ => by simp [sign, Int.mul_one] + | _, 0 => by simp [sign, Int.mul_zero] + | _, -[_+1] => by simp [sign, Int.mul_neg, Int.mul_one] + +/-! ### `/` and ordering -/ + +protected theorem ediv_mul_le (a : Int) {b : Int} (H : b ≠ 0) : a / b * b ≤ a := + Int.le_of_sub_nonneg <| by rw [Int.mul_comm, ← emod_def]; apply emod_nonneg _ H + +theorem le_of_mul_le_mul_left {a b c : Int} (w : a * b ≤ a * c) (h : 0 < a) : b ≤ c := by + have w := Int.sub_nonneg_of_le w + rw [← Int.mul_sub] at w + have w := Int.ediv_nonneg w (Int.le_of_lt h) + rw [Int.mul_ediv_cancel_left _ (Int.ne_of_gt h)] at w + exact Int.le_of_sub_nonneg w + +theorem le_of_mul_le_mul_right {a b c : Int} (w : b * a ≤ c * a) (h : 0 < a) : b ≤ c := by + rw [Int.mul_comm b, Int.mul_comm c] at w + exact le_of_mul_le_mul_left w h + +protected theorem ediv_le_of_le_mul {a b c : Int} (H : 0 < c) (H' : a ≤ b * c) : a / c ≤ b := + le_of_mul_le_mul_right (Int.le_trans (Int.ediv_mul_le _ (Int.ne_of_gt H)) H') H + +protected theorem mul_lt_of_lt_ediv {a b c : Int} (H : 0 < c) (H3 : a < b / c) : a * c < b := + Int.lt_of_not_ge <| mt (Int.ediv_le_of_le_mul H) (Int.not_le_of_gt H3) + +protected theorem mul_le_of_le_ediv {a b c : Int} (H1 : 0 < c) (H2 : a ≤ b / c) : a * c ≤ b := + Int.le_trans (Int.mul_le_mul_of_nonneg_right H2 (Int.le_of_lt H1)) + (Int.ediv_mul_le _ (Int.ne_of_gt H1)) + +protected theorem ediv_lt_of_lt_mul {a b c : Int} (H : 0 < c) (H' : a < b * c) : a / c < b := + Int.lt_of_not_ge <| mt (Int.mul_le_of_le_ediv H) (Int.not_le_of_gt H') + +theorem lt_of_mul_lt_mul_left {a b c : Int} (w : a * b < a * c) (h : 0 ≤ a) : b < c := by + rcases Int.lt_trichotomy b c with lt | rfl | gt + · exact lt + · exact False.elim (Int.lt_irrefl _ w) + · rcases Int.lt_trichotomy a 0 with a_lt | rfl | a_gt + · exact False.elim (Int.lt_irrefl _ (Int.lt_of_lt_of_le a_lt h)) + · exact False.elim (Int.lt_irrefl b (by simp at w)) + · have := le_of_mul_le_mul_left (Int.le_of_lt w) a_gt + exact False.elim (Int.lt_irrefl _ (Int.lt_of_lt_of_le gt this)) + +theorem lt_of_mul_lt_mul_right {a b c : Int} (w : b * a < c * a) (h : 0 ≤ a) : b < c := by + rw [Int.mul_comm b, Int.mul_comm c] at w + exact lt_of_mul_lt_mul_left w h + +protected theorem le_ediv_of_mul_le {a b c : Int} (H1 : 0 < c) (H2 : a * c ≤ b) : a ≤ b / c := + le_of_lt_add_one <| + lt_of_mul_lt_mul_right (Int.lt_of_le_of_lt H2 (lt_ediv_add_one_mul_self _ H1)) (Int.le_of_lt H1) + +protected theorem le_ediv_iff_mul_le {a b c : Int} (H : 0 < c) : a ≤ b / c ↔ a * c ≤ b := + ⟨Int.mul_le_of_le_ediv H, Int.le_ediv_of_mul_le H⟩ + +protected theorem ediv_le_ediv {a b c : Int} (H : 0 < c) (H' : a ≤ b) : a / c ≤ b / c := + Int.le_ediv_of_mul_le H (Int.le_trans (Int.ediv_mul_le _ (Int.ne_of_gt H)) H') + +protected theorem lt_mul_of_ediv_lt {a b c : Int} (H1 : 0 < c) (H2 : a / c < b) : a < b * c := + Int.lt_of_not_ge <| mt (Int.le_ediv_of_mul_le H1) (Int.not_le_of_gt H2) + +protected theorem ediv_lt_iff_lt_mul {a b c : Int} (H : 0 < c) : a / c < b ↔ a < b * c := + ⟨Int.lt_mul_of_ediv_lt H, Int.ediv_lt_of_lt_mul H⟩ + +protected theorem le_mul_of_ediv_le {a b c : Int} (H1 : 0 ≤ b) (H2 : b ∣ a) (H3 : a / b ≤ c) : + a ≤ c * b := by + rw [← Int.ediv_mul_cancel H2]; exact Int.mul_le_mul_of_nonneg_right H3 H1 + +protected theorem lt_ediv_of_mul_lt {a b c : Int} (H1 : 0 ≤ b) (H2 : b ∣ c) (H3 : a * b < c) : + a < c / b := + Int.lt_of_not_ge <| mt (Int.le_mul_of_ediv_le H1 H2) (Int.not_le_of_gt H3) + +protected theorem lt_ediv_iff_mul_lt {a b : Int} (c : Int) (H : 0 < c) (H' : c ∣ b) : + a < b / c ↔ a * c < b := + ⟨Int.mul_lt_of_lt_ediv H, Int.lt_ediv_of_mul_lt (Int.le_of_lt H) H'⟩ + +theorem ediv_pos_of_pos_of_dvd {a b : Int} (H1 : 0 < a) (H2 : 0 ≤ b) (H3 : b ∣ a) : 0 < a / b := + Int.lt_ediv_of_mul_lt H2 H3 (by rwa [Int.zero_mul]) + +theorem ediv_eq_ediv_of_mul_eq_mul {a b c d : Int} + (H2 : d ∣ c) (H3 : b ≠ 0) (H4 : d ≠ 0) (H5 : a * d = b * c) : a / b = c / d := + Int.ediv_eq_of_eq_mul_right H3 <| by + rw [← Int.mul_ediv_assoc _ H2]; exact (Int.ediv_eq_of_eq_mul_left H4 H5.symm).symm + +/-! ### div -/ + +@[simp] protected theorem div_one : ∀ a : Int, a.div 1 = a + | (n:Nat) => congrArg ofNat (Nat.div_one _) + | -[n+1] => by simp [Int.div, neg_ofNat_succ] + +@[simp] protected theorem div_neg : ∀ a b : Int, a.div (-b) = -(a.div b) + | ofNat m, 0 => show ofNat (m / 0) = -↑(m / 0) by rw [Nat.div_zero]; rfl + | ofNat m, -[n+1] | -[m+1], succ n => (Int.neg_neg _).symm + | ofNat m, succ n | -[m+1], 0 | -[m+1], -[n+1] => rfl + +@[simp] protected theorem neg_div : ∀ a b : Int, (-a).div b = -(a.div b) + | 0, n => by simp [Int.neg_zero] + | succ m, (n:Nat) | -[m+1], 0 | -[m+1], -[n+1] => rfl + | succ m, -[n+1] | -[m+1], succ n => (Int.neg_neg _).symm + +protected theorem neg_div_neg (a b : Int) : (-a).div (-b) = a.div b := by + simp [Int.div_neg, Int.neg_div, Int.neg_neg] + +protected theorem div_nonneg {a b : Int} (Ha : 0 ≤ a) (Hb : 0 ≤ b) : 0 ≤ a.div b := + match a, b, eq_ofNat_of_zero_le Ha, eq_ofNat_of_zero_le Hb with + | _, _, ⟨_, rfl⟩, ⟨_, rfl⟩ => ofNat_zero_le _ + +protected theorem div_nonpos {a b : Int} (Ha : 0 ≤ a) (Hb : b ≤ 0) : a.div b ≤ 0 := + Int.nonpos_of_neg_nonneg <| Int.div_neg .. ▸ Int.div_nonneg Ha (Int.neg_nonneg_of_nonpos Hb) + +theorem div_eq_zero_of_lt {a b : Int} (H1 : 0 ≤ a) (H2 : a < b) : a.div b = 0 := + match a, b, eq_ofNat_of_zero_le H1, eq_succ_of_zero_lt (Int.lt_of_le_of_lt H1 H2) with + | _, _, ⟨_, rfl⟩, ⟨_, rfl⟩ => congrArg Nat.cast <| Nat.div_eq_of_lt <| ofNat_lt.1 H2 + +@[simp] protected theorem mul_div_cancel (a : Int) {b : Int} (H : b ≠ 0) : (a * b).div b = a := + have : ∀ {a b : Nat}, (b : Int) ≠ 0 → (div (a * b) b : Int) = a := fun H => by + rw [← ofNat_mul, ← ofNat_div, + Nat.mul_div_cancel _ <| Nat.pos_of_ne_zero <| Int.ofNat_ne_zero.1 H] + match a, b, a.eq_nat_or_neg, b.eq_nat_or_neg with + | _, _, ⟨a, .inl rfl⟩, ⟨b, .inl rfl⟩ => this H + | _, _, ⟨a, .inl rfl⟩, ⟨b, .inr rfl⟩ => by + rw [Int.mul_neg, Int.neg_div, Int.div_neg, Int.neg_neg, + this (Int.neg_ne_zero.1 H)] + | _, _, ⟨a, .inr rfl⟩, ⟨b, .inl rfl⟩ => by rw [Int.neg_mul, Int.neg_div, this H] + | _, _, ⟨a, .inr rfl⟩, ⟨b, .inr rfl⟩ => by + rw [Int.neg_mul_neg, Int.div_neg, this (Int.neg_ne_zero.1 H)] + +@[simp] protected theorem mul_div_cancel_left (b : Int) (H : a ≠ 0) : (a * b).div a = b := + Int.mul_comm .. ▸ Int.mul_div_cancel _ H + +@[simp] protected theorem div_self {a : Int} (H : a ≠ 0) : a.div a = 1 := by + have := Int.mul_div_cancel 1 H; rwa [Int.one_mul] at this + +theorem mul_div_cancel_of_mod_eq_zero {a b : Int} (H : a.mod b = 0) : b * (a.div b) = a := by + have := mod_add_div a b; rwa [H, Int.zero_add] at this + +theorem div_mul_cancel_of_mod_eq_zero {a b : Int} (H : a.mod b = 0) : a.div b * b = a := by + rw [Int.mul_comm, mul_div_cancel_of_mod_eq_zero H] + +theorem dvd_of_mod_eq_zero {a b : Int} (H : mod b a = 0) : a ∣ b := + ⟨b.div a, (mul_div_cancel_of_mod_eq_zero H).symm⟩ + +protected theorem mul_div_assoc (a : Int) : ∀ {b c : Int}, c ∣ b → (a * b).div c = a * (b.div c) + | _, c, ⟨d, rfl⟩ => + if cz : c = 0 then by simp [cz, Int.mul_zero] else by + rw [Int.mul_left_comm, Int.mul_div_cancel_left _ cz, Int.mul_div_cancel_left _ cz] + +protected theorem mul_div_assoc' (b : Int) {a c : Int} (h : c ∣ a) : + (a * b).div c = a.div c * b := by + rw [Int.mul_comm, Int.mul_div_assoc _ h, Int.mul_comm] + +theorem div_dvd_div : ∀ {a b c : Int}, a ∣ b → b ∣ c → b.div a ∣ c.div a + | a, _, _, ⟨b, rfl⟩, ⟨c, rfl⟩ => by + by_cases az : a = 0 + · simp [az] + · rw [Int.mul_div_cancel_left _ az, Int.mul_assoc, Int.mul_div_cancel_left _ az] + apply Int.dvd_mul_right + +@[simp] theorem natAbs_div (a b : Int) : natAbs (a.div b) = (natAbs a).div (natAbs b) := + match a, b, eq_nat_or_neg a, eq_nat_or_neg b with + | _, _, ⟨_, .inl rfl⟩, ⟨_, .inl rfl⟩ => rfl + | _, _, ⟨_, .inl rfl⟩, ⟨_, .inr rfl⟩ => by rw [Int.div_neg, natAbs_neg, natAbs_neg]; rfl + | _, _, ⟨_, .inr rfl⟩, ⟨_, .inl rfl⟩ => by rw [Int.neg_div, natAbs_neg, natAbs_neg]; rfl + | _, _, ⟨_, .inr rfl⟩, ⟨_, .inr rfl⟩ => by rw [Int.neg_div_neg, natAbs_neg, natAbs_neg]; rfl + +protected theorem div_eq_of_eq_mul_right {a b c : Int} + (H1 : b ≠ 0) (H2 : a = b * c) : a.div b = c := by rw [H2, Int.mul_div_cancel_left _ H1] + +protected theorem eq_div_of_mul_eq_right {a b c : Int} + (H1 : a ≠ 0) (H2 : a * b = c) : b = c.div a := + (Int.div_eq_of_eq_mul_right H1 H2.symm).symm + +/-! ### (t-)mod -/ + +theorem ofNat_mod (m n : Nat) : (↑(m % n) : Int) = mod m n := rfl + +@[simp] theorem mod_one (a : Int) : mod a 1 = 0 := by + simp [mod_def, Int.div_one, Int.one_mul, Int.sub_self] + +theorem mod_eq_of_lt {a b : Int} (H1 : 0 ≤ a) (H2 : a < b) : mod a b = a := by + rw [mod_eq_emod H1 (Int.le_trans H1 (Int.le_of_lt H2)), emod_eq_of_lt H1 H2] + +theorem mod_lt_of_pos (a : Int) {b : Int} (H : 0 < b) : mod a b < b := + match a, b, eq_succ_of_zero_lt H with + | ofNat _, _, ⟨n, rfl⟩ => ofNat_lt.2 <| Nat.mod_lt _ n.succ_pos + | -[_+1], _, ⟨n, rfl⟩ => Int.lt_of_le_of_lt + (Int.neg_nonpos_of_nonneg <| Int.ofNat_nonneg _) (ofNat_pos.2 n.succ_pos) + +theorem mod_nonneg : ∀ {a : Int} (b : Int), 0 ≤ a → 0 ≤ mod a b + | ofNat _, -[_+1], _ | ofNat _, ofNat _, _ => ofNat_nonneg _ + +@[simp] theorem mod_neg (a b : Int) : mod a (-b) = mod a b := by + rw [mod_def, mod_def, Int.div_neg, Int.neg_mul_neg] + +@[simp] theorem mul_mod_left (a b : Int) : (a * b).mod b = 0 := + if h : b = 0 then by simp [h, Int.mul_zero] else by + rw [Int.mod_def, Int.mul_div_cancel _ h, Int.mul_comm, Int.sub_self] + +@[simp] theorem mul_mod_right (a b : Int) : (a * b).mod a = 0 := by + rw [Int.mul_comm, mul_mod_left] + +theorem mod_eq_zero_of_dvd : ∀ {a b : Int}, a ∣ b → mod b a = 0 + | _, _, ⟨_, rfl⟩ => mul_mod_right .. + +theorem dvd_iff_mod_eq_zero (a b : Int) : a ∣ b ↔ mod b a = 0 := + ⟨mod_eq_zero_of_dvd, dvd_of_mod_eq_zero⟩ + +protected theorem div_mul_cancel {a b : Int} (H : b ∣ a) : a.div b * b = a := + div_mul_cancel_of_mod_eq_zero (mod_eq_zero_of_dvd H) + +protected theorem mul_div_cancel' {a b : Int} (H : a ∣ b) : a * b.div a = b := by + rw [Int.mul_comm, Int.div_mul_cancel H] + +protected theorem eq_mul_of_div_eq_right {a b c : Int} + (H1 : b ∣ a) (H2 : a.div b = c) : a = b * c := by rw [← H2, Int.mul_div_cancel' H1] + +@[simp] theorem mod_self {a : Int} : a.mod a = 0 := by + have := mul_mod_left 1 a; rwa [Int.one_mul] at this + +theorem lt_div_add_one_mul_self (a : Int) {b : Int} (H : 0 < b) : a < (a.div b + 1) * b := by + rw [Int.add_mul, Int.one_mul, Int.mul_comm] + exact Int.lt_add_of_sub_left_lt <| Int.mod_def .. ▸ mod_lt_of_pos _ H + +protected theorem div_eq_iff_eq_mul_right {a b c : Int} + (H : b ≠ 0) (H' : b ∣ a) : a.div b = c ↔ a = b * c := + ⟨Int.eq_mul_of_div_eq_right H', Int.div_eq_of_eq_mul_right H⟩ + +protected theorem div_eq_iff_eq_mul_left {a b c : Int} + (H : b ≠ 0) (H' : b ∣ a) : a.div b = c ↔ a = c * b := by + rw [Int.mul_comm]; exact Int.div_eq_iff_eq_mul_right H H' + +protected theorem eq_mul_of_div_eq_left {a b c : Int} + (H1 : b ∣ a) (H2 : a.div b = c) : a = c * b := by + rw [Int.mul_comm, Int.eq_mul_of_div_eq_right H1 H2] + +protected theorem div_eq_of_eq_mul_left {a b c : Int} + (H1 : b ≠ 0) (H2 : a = c * b) : a.div b = c := + Int.div_eq_of_eq_mul_right H1 (by rw [Int.mul_comm, H2]) + +protected theorem eq_zero_of_div_eq_zero {d n : Int} (h : d ∣ n) (H : n.div d = 0) : n = 0 := by + rw [← Int.mul_div_cancel' h, H, Int.mul_zero] + +@[simp] protected theorem div_left_inj {a b d : Int} + (hda : d ∣ a) (hdb : d ∣ b) : a.div d = b.div d ↔ a = b := by + refine ⟨fun h => ?_, congrArg (div · d)⟩ + rw [← Int.mul_div_cancel' hda, ← Int.mul_div_cancel' hdb, h] + +theorem div_sign : ∀ a b, a.div (sign b) = a * sign b + | _, succ _ => by simp [sign, Int.mul_one] + | _, 0 => by simp [sign, Int.mul_zero] + | _, -[_+1] => by simp [sign, Int.mul_neg, Int.mul_one] + +protected theorem sign_eq_div_abs (a : Int) : sign a = a.div (natAbs a) := + if az : a = 0 then by simp [az] else + (Int.div_eq_of_eq_mul_left (ofNat_ne_zero.2 <| natAbs_ne_zero.2 az) + (sign_mul_natAbs _).symm).symm + +/-! ### fdiv -/ + +theorem fdiv_nonneg {a b : Int} (Ha : 0 ≤ a) (Hb : 0 ≤ b) : 0 ≤ a.fdiv b := + match a, b, eq_ofNat_of_zero_le Ha, eq_ofNat_of_zero_le Hb with + | _, _, ⟨_, rfl⟩, ⟨_, rfl⟩ => ofNat_fdiv .. ▸ ofNat_zero_le _ + +theorem fdiv_nonpos : ∀ {a b : Int}, 0 ≤ a → b ≤ 0 → a.fdiv b ≤ 0 + | 0, 0, _, _ | 0, -[_+1], _, _ | succ _, 0, _, _ | succ _, -[_+1], _, _ => ⟨_⟩ + +theorem fdiv_neg' : ∀ {a b : Int}, a < 0 → 0 < b → a.fdiv b < 0 + | -[_+1], succ _, _, _ => negSucc_lt_zero _ + +@[simp] theorem fdiv_one : ∀ a : Int, a.fdiv 1 = a + | 0 => rfl + | succ _ => congrArg Nat.cast (Nat.div_one _) + | -[_+1] => congrArg negSucc (Nat.div_one _) + +@[simp] theorem mul_fdiv_cancel (a : Int) {b : Int} (H : b ≠ 0) : fdiv (a * b) b = a := + if b0 : 0 ≤ b then by + rw [fdiv_eq_ediv _ b0, mul_ediv_cancel _ H] + else + match a, b, Int.not_le.1 b0 with + | 0, _, _ => by simp [Int.zero_mul] + | succ a, -[b+1], _ => congrArg ofNat <| Nat.mul_div_cancel (succ a) b.succ_pos + | -[a+1], -[b+1], _ => congrArg negSucc <| Nat.div_eq_of_lt_le + (Nat.le_of_lt_succ <| Nat.mul_lt_mul_of_pos_right a.lt_succ_self b.succ_pos) + (Nat.lt_succ_self _) + +@[simp] theorem mul_fdiv_cancel_left (b : Int) (H : a ≠ 0) : fdiv (a * b) a = b := + Int.mul_comm .. ▸ Int.mul_fdiv_cancel _ H + +@[simp] protected theorem fdiv_self {a : Int} (H : a ≠ 0) : a.fdiv a = 1 := by + have := Int.mul_fdiv_cancel 1 H; rwa [Int.one_mul] at this + +theorem lt_fdiv_add_one_mul_self (a : Int) {b : Int} (H : 0 < b) : a < (a.fdiv b + 1) * b := + Int.fdiv_eq_ediv _ (Int.le_of_lt H) ▸ lt_ediv_add_one_mul_self a H + +/-! ### fmod -/ + +theorem ofNat_fmod (m n : Nat) : ↑(m % n) = fmod m n := by + cases m <;> simp [fmod, Nat.succ_eq_add_one] + +@[simp] theorem fmod_one (a : Int) : a.fmod 1 = 0 := by + simp [fmod_def, Int.one_mul, Int.sub_self] + +theorem fmod_eq_of_lt {a b : Int} (H1 : 0 ≤ a) (H2 : a < b) : a.fmod b = a := by + rw [fmod_eq_emod _ (Int.le_trans H1 (Int.le_of_lt H2)), emod_eq_of_lt H1 H2] + +theorem fmod_nonneg {a b : Int} (ha : 0 ≤ a) (hb : 0 ≤ b) : 0 ≤ a.fmod b := + fmod_eq_mod ha hb ▸ mod_nonneg _ ha + +theorem fmod_nonneg' (a : Int) {b : Int} (hb : 0 < b) : 0 ≤ a.fmod b := + fmod_eq_emod _ (Int.le_of_lt hb) ▸ emod_nonneg _ (Int.ne_of_lt hb).symm + +theorem fmod_lt_of_pos (a : Int) {b : Int} (H : 0 < b) : a.fmod b < b := + fmod_eq_emod _ (Int.le_of_lt H) ▸ emod_lt_of_pos a H + +@[simp] theorem mul_fmod_left (a b : Int) : (a * b).fmod b = 0 := + if h : b = 0 then by simp [h, Int.mul_zero] else by + rw [Int.fmod_def, Int.mul_fdiv_cancel _ h, Int.mul_comm, Int.sub_self] + +@[simp] theorem mul_fmod_right (a b : Int) : (a * b).fmod a = 0 := by + rw [Int.mul_comm, mul_fmod_left] + +@[simp] theorem fmod_self {a : Int} : a.fmod a = 0 := by + have := mul_fmod_left 1 a; rwa [Int.one_mul] at this + +/-! ### Theorems crossing div/mod versions -/ + +theorem div_eq_ediv_of_dvd {a b : Int} (h : b ∣ a) : a.div b = a / b := by + by_cases b0 : b = 0 + · simp [b0] + · rw [Int.div_eq_iff_eq_mul_left b0 h, ← Int.ediv_eq_iff_eq_mul_left b0 h] + +theorem fdiv_eq_ediv_of_dvd : ∀ {a b : Int}, b ∣ a → a.fdiv b = a / b + | _, b, ⟨c, rfl⟩ => by + by_cases bz : b = 0 + · simp [bz] + · rw [mul_fdiv_cancel_left _ bz, mul_ediv_cancel_left _ bz] + +/-! ### bmod -/ @[simp] theorem bmod_emod : bmod x m % m = x % m := by dsimp [bmod] @@ -397,3 +1066,131 @@ theorem bmod_add_bmod_congr : Int.bmod (Int.bmod x n + y) n = Int.bmod (x + y) n @[simp] theorem add_bmod_bmod : Int.bmod (x + Int.bmod y n) n = Int.bmod (x + y) n := by rw [Int.add_comm x, Int.bmod_add_bmod_congr, Int.add_comm y] + +theorem emod_bmod {x : Int} {m : Nat} : bmod (x % m) m = bmod x m := by + simp [bmod] + +@[simp] theorem bmod_bmod : bmod (bmod x m) m = bmod x m := by + rw [bmod, bmod_emod] + rfl + +@[simp] theorem bmod_zero : Int.bmod 0 m = 0 := by + dsimp [bmod] + simp only [zero_emod, Int.zero_sub, ite_eq_left_iff, Int.neg_eq_zero] + intro h + rw [@Int.not_lt] at h + match m with + | 0 => rfl + | (m+1) => + exfalso + rw [natCast_add, ofNat_one, Int.add_assoc, add_ediv_of_dvd_right] at h + change _ + 2 / 2 ≤ 0 at h + rw [Int.ediv_self, ← ofNat_two, ← ofNat_ediv, add_one_le_iff, ← @Int.not_le] at h + exact h (ofNat_nonneg _) + all_goals decide + +theorem dvd_bmod_sub_self {x : Int} {m : Nat} : (m : Int) ∣ bmod x m - x := by + dsimp [bmod] + split + · exact dvd_emod_sub_self + · rw [Int.sub_sub, Int.add_comm, ← Int.sub_sub] + exact Int.dvd_sub dvd_emod_sub_self (Int.dvd_refl _) + +theorem le_bmod {x : Int} {m : Nat} (h : 0 < m) : - (m/2) ≤ Int.bmod x m := by + dsimp [bmod] + have v : (m : Int) % 2 = 0 ∨ (m : Int) % 2 = 1 := emod_two_eq _ + split <;> rename_i w + · refine Int.le_trans ?_ (Int.emod_nonneg _ ?_) + · exact Int.neg_nonpos_of_nonneg (Int.ediv_nonneg (Int.ofNat_nonneg _) (by decide)) + · exact Int.ne_of_gt (ofNat_pos.mpr h) + · simp [Int.not_lt] at w + refine Int.le_trans ?_ (Int.sub_le_sub_right w _) + rw [← ediv_add_emod m 2] + generalize (m : Int) / 2 = q + generalize h : (m : Int) % 2 = r at * + rcases v with rfl | rfl + · rw [Int.add_zero, Int.mul_ediv_cancel_left, Int.add_ediv_of_dvd_left, + Int.mul_ediv_cancel_left, show (1 / 2 : Int) = 0 by decide, Int.add_zero, + Int.neg_eq_neg_one_mul] + conv => rhs; congr; rw [← Int.one_mul q] + rw [← Int.sub_mul, show (1 - 2 : Int) = -1 by decide] + apply Int.le_refl + all_goals try decide + all_goals apply Int.dvd_mul_right + · rw [Int.add_ediv_of_dvd_left, Int.mul_ediv_cancel_left, + show (1 / 2 : Int) = 0 by decide, Int.add_assoc, Int.add_ediv_of_dvd_left, + Int.mul_ediv_cancel_left, show ((1 + 1) / 2 : Int) = 1 by decide, ← Int.sub_sub, + Int.sub_eq_add_neg, Int.sub_eq_add_neg, Int.add_right_comm, Int.add_assoc q, + show (1 + -1 : Int) = 0 by decide, Int.add_zero, ← Int.neg_mul] + rw [Int.neg_eq_neg_one_mul] + conv => rhs; congr; rw [← Int.one_mul q] + rw [← Int.add_mul, show (1 + -2 : Int) = -1 by decide] + apply Int.le_refl + all_goals try decide + all_goals try apply Int.dvd_mul_right + +theorem bmod_lt {x : Int} {m : Nat} (h : 0 < m) : bmod x m < (m + 1) / 2 := by + dsimp [bmod] + split + · assumption + · apply Int.lt_of_lt_of_le + · show _ < 0 + have : x % m < m := emod_lt_of_pos x (ofNat_pos.mpr h) + exact Int.sub_neg_of_lt this + · exact Int.le.intro_sub _ rfl + +theorem bmod_le {x : Int} {m : Nat} (h : 0 < m) : bmod x m ≤ (m - 1) / 2 := by + refine lt_add_one_iff.mp ?_ + calc + bmod x m < (m + 1) / 2 := bmod_lt h + _ = ((m + 1 - 2) + 2)/2 := by simp + _ = (m - 1) / 2 + 1 := by + rw [add_ediv_of_dvd_right] + · simp (config := {decide := true}) only [Int.ediv_self] + congr 2 + rw [Int.add_sub_assoc, ← Int.sub_neg] + congr + · trivial + +-- This could be strengthed by changing to `w : x ≠ -1` if needed. +theorem bmod_natAbs_plus_one (x : Int) (w : 1 < x.natAbs) : bmod x (x.natAbs + 1) = - x.sign := by + have t₁ : ∀ (x : Nat), x % (x + 2) = x := + fun x => Nat.mod_eq_of_lt (Nat.lt_succ_of_lt (Nat.lt.base x)) + have t₂ : ∀ (x : Int), 0 ≤ x → x % (x + 2) = x := fun x h => by + match x, h with + | Int.ofNat x, _ => erw [← Int.ofNat_two, ← ofNat_add, ← ofNat_emod, t₁]; rfl + cases x with + | ofNat x => + simp only [bmod, ofNat_eq_coe, natAbs_ofNat, natCast_add, ofNat_one, + emod_self_add_one (ofNat_nonneg x)] + match x with + | 0 => rw [if_pos] <;> simp (config := {decide := true}) + | (x+1) => + rw [if_neg] + · simp [← Int.sub_sub] + · refine Int.not_lt.mpr ?_ + simp only [← natCast_add, ← ofNat_one, ← ofNat_two, ← ofNat_ediv] + match x with + | 0 => apply Int.le_refl + | (x+1) => + refine Int.ofNat_le.mpr ?_ + apply Nat.div_le_of_le_mul + simp only [Nat.two_mul, Nat.add_assoc] + apply Nat.add_le_add_left (Nat.add_le_add_left (Nat.add_le_add_left (Nat.le_add_left + _ _) _) _) + | negSucc x => + rw [bmod, natAbs_negSucc, natCast_add, ofNat_one, sign_negSucc, Int.neg_neg, + Nat.succ_eq_add_one, negSucc_emod] + erw [t₂] + · rw [natCast_add, ofNat_one, Int.add_sub_cancel, Int.add_comm, Int.add_sub_cancel, if_pos] + · match x, w with + | (x+1), _ => + rw [Int.add_assoc, add_ediv_of_dvd_right, show (1 + 1 : Int) = 2 by decide, Int.ediv_self] + apply Int.lt_add_one_of_le + rw [Int.add_comm, ofNat_add, Int.add_assoc, add_ediv_of_dvd_right, + show ((1 : Nat) + 1 : Int) = 2 by decide, Int.ediv_self] + apply Int.le_add_of_nonneg_left + exact Int.le.intro_sub _ rfl + all_goals decide + · exact ofNat_nonneg x + · exact succ_ofNat_pos (x + 1) diff --git a/src/Init/Data/Int/Gcd.lean b/src/Init/Data/Int/Gcd.lean index abb5234c7e04..3e8748123f6a 100644 --- a/src/Init/Data/Int/Gcd.lean +++ b/src/Init/Data/Int/Gcd.lean @@ -6,7 +6,12 @@ Authors: Mario Carneiro prelude import Init.Data.Int.Basic import Init.Data.Nat.Gcd +import Init.Data.Nat.Lcm +import Init.Data.Int.DivModLemmas +/-! +Definition and lemmas for gcd and lcm over Int +-/ namespace Int /-! ## gcd -/ @@ -14,4 +19,37 @@ namespace Int /-- Computes the greatest common divisor of two integers, as a `Nat`. -/ def gcd (m n : Int) : Nat := m.natAbs.gcd n.natAbs +theorem gcd_dvd_left {a b : Int} : (gcd a b : Int) ∣ a := by + have := Nat.gcd_dvd_left a.natAbs b.natAbs + rw [← Int.ofNat_dvd] at this + exact Int.dvd_trans this natAbs_dvd_self + +theorem gcd_dvd_right {a b : Int} : (gcd a b : Int) ∣ b := by + have := Nat.gcd_dvd_right a.natAbs b.natAbs + rw [← Int.ofNat_dvd] at this + exact Int.dvd_trans this natAbs_dvd_self + +@[simp] theorem one_gcd {a : Int} : gcd 1 a = 1 := by simp [gcd] +@[simp] theorem gcd_one {a : Int} : gcd a 1 = 1 := by simp [gcd] + +@[simp] theorem neg_gcd {a b : Int} : gcd (-a) b = gcd a b := by simp [gcd] +@[simp] theorem gcd_neg {a b : Int} : gcd a (-b) = gcd a b := by simp [gcd] + +/-! ## lcm -/ + +/-- Computes the least common multiple of two integers, as a `Nat`. -/ +def lcm (m n : Int) : Nat := m.natAbs.lcm n.natAbs + +theorem lcm_ne_zero (hm : m ≠ 0) (hn : n ≠ 0) : lcm m n ≠ 0 := by + simp only [lcm] + apply Nat.lcm_ne_zero <;> simpa + +theorem dvd_lcm_left {a b : Int} : a ∣ lcm a b := + Int.dvd_trans dvd_natAbs_self (Int.ofNat_dvd.mpr (Nat.dvd_lcm_left a.natAbs b.natAbs)) + +theorem dvd_lcm_right {a b : Int} : b ∣ lcm a b := + Int.dvd_trans dvd_natAbs_self (Int.ofNat_dvd.mpr (Nat.dvd_lcm_right a.natAbs b.natAbs)) + +@[simp] theorem lcm_self {a : Int} : lcm a a = a.natAbs := Nat.lcm_self _ + end Int diff --git a/src/Init/Data/Int/Lemmas.lean b/src/Init/Data/Int/Lemmas.lean index 82ce1f69f2a4..b0c98c77ca7c 100644 --- a/src/Init/Data/Int/Lemmas.lean +++ b/src/Init/Data/Int/Lemmas.lean @@ -499,33 +499,6 @@ theorem eq_one_of_mul_eq_self_left {a b : Int} (Hpos : a ≠ 0) (H : b * a = a) theorem eq_one_of_mul_eq_self_right {a b : Int} (Hpos : b ≠ 0) (H : b * a = b) : a = 1 := Int.eq_of_mul_eq_mul_left Hpos <| by rw [Int.mul_one, H] -/-! # pow -/ - -protected theorem pow_zero (b : Int) : b^0 = 1 := rfl - -protected theorem pow_succ (b : Int) (e : Nat) : b ^ (e+1) = (b ^ e) * b := rfl -protected theorem pow_succ' (b : Int) (e : Nat) : b ^ (e+1) = b * (b ^ e) := by - rw [Int.mul_comm, Int.pow_succ] - -theorem pow_le_pow_of_le_left {n m : Nat} (h : n ≤ m) : ∀ (i : Nat), n^i ≤ m^i - | 0 => Nat.le_refl _ - | succ i => Nat.mul_le_mul (pow_le_pow_of_le_left h i) h - -theorem pow_le_pow_of_le_right {n : Nat} (hx : n > 0) {i : Nat} : ∀ {j}, i ≤ j → n^i ≤ n^j - | 0, h => - have : i = 0 := eq_zero_of_le_zero h - this.symm ▸ Nat.le_refl _ - | succ j, h => - match le_or_eq_of_le_succ h with - | Or.inl h => show n^i ≤ n^j * n from - have : n^i * 1 ≤ n^j * n := Nat.mul_le_mul (pow_le_pow_of_le_right hx h) hx - Nat.mul_one (n^i) ▸ this - | Or.inr h => - h.symm ▸ Nat.le_refl _ - -theorem pos_pow_of_pos {n : Nat} (m : Nat) (h : 0 < n) : 0 < n^m := - pow_le_pow_of_le_right h (Nat.zero_le _) - /-! NatCast lemmas -/ /-! @@ -545,10 +518,4 @@ theorem natCast_one : ((1 : Nat) : Int) = (1 : Int) := rfl @[simp] theorem natCast_mul (a b : Nat) : ((a * b : Nat) : Int) = (a : Int) * (b : Int) := by simp -theorem natCast_pow (b n : Nat) : ((b^n : Nat) : Int) = (b : Int) ^ n := by - match n with - | 0 => rfl - | n + 1 => - simp only [Nat.pow_succ, Int.pow_succ, natCast_mul, natCast_pow _ n] - end Int diff --git a/src/Init/Data/Int/Order.lean b/src/Init/Data/Int/Order.lean index bf99711c2db8..4e59cc3f1d19 100644 --- a/src/Init/Data/Int/Order.lean +++ b/src/Init/Data/Int/Order.lean @@ -6,6 +6,7 @@ Authors: Jeremy Avigad, Deniz Aydin, Floris van Doorn, Mario Carneiro prelude import Init.Data.Int.Lemmas import Init.ByCases +import Init.RCases /-! # Results about the order properties of the integers, and the integers as an ordered ring. @@ -498,3 +499,524 @@ theorem toNat_add_nat {a : Int} (ha : 0 ≤ a) (n : Nat) : (a + n).toNat = a.toN @[simp] theorem toNat_neg_nat : ∀ n : Nat, (-(n : Int)).toNat = 0 | 0 => rfl | _+1 => rfl + +/-! ### toNat' -/ + +theorem mem_toNat' : ∀ (a : Int) (n : Nat), toNat' a = some n ↔ a = n + | (m : Nat), n => by simp [toNat', Int.ofNat_inj] + | -[m+1], n => by constructor <;> nofun + +/-! ## Order properties of the integers -/ + +protected theorem lt_of_not_ge {a b : Int} : ¬a ≤ b → b < a := Int.not_le.mp +protected theorem not_le_of_gt {a b : Int} : b < a → ¬a ≤ b := Int.not_le.mpr + +protected theorem le_of_not_le {a b : Int} : ¬ a ≤ b → b ≤ a := (Int.le_total a b).resolve_left + +@[simp] theorem negSucc_not_pos (n : Nat) : 0 < -[n+1] ↔ False := by + simp only [Int.not_lt, iff_false]; constructor + +theorem eq_negSucc_of_lt_zero : ∀ {a : Int}, a < 0 → ∃ n : Nat, a = -[n+1] + | ofNat _, h => absurd h (Int.not_lt.2 (ofNat_zero_le _)) + | -[n+1], _ => ⟨n, rfl⟩ + +protected theorem lt_of_add_lt_add_left {a b c : Int} (h : a + b < a + c) : b < c := by + have : -a + (a + b) < -a + (a + c) := Int.add_lt_add_left h _ + simp [Int.neg_add_cancel_left] at this + assumption + +protected theorem lt_of_add_lt_add_right {a b c : Int} (h : a + b < c + b) : a < c := + Int.lt_of_add_lt_add_left (a := b) <| by rwa [Int.add_comm b a, Int.add_comm b c] + +protected theorem add_lt_add_iff_left (a : Int) : a + b < a + c ↔ b < c := + ⟨Int.lt_of_add_lt_add_left, (Int.add_lt_add_left · _)⟩ + +protected theorem add_lt_add_iff_right (c : Int) : a + c < b + c ↔ a < b := + ⟨Int.lt_of_add_lt_add_right, (Int.add_lt_add_right · _)⟩ + +protected theorem add_lt_add {a b c d : Int} (h₁ : a < b) (h₂ : c < d) : a + c < b + d := + Int.lt_trans (Int.add_lt_add_right h₁ c) (Int.add_lt_add_left h₂ b) + +protected theorem add_lt_add_of_le_of_lt {a b c d : Int} (h₁ : a ≤ b) (h₂ : c < d) : + a + c < b + d := + Int.lt_of_le_of_lt (Int.add_le_add_right h₁ c) (Int.add_lt_add_left h₂ b) + +protected theorem add_lt_add_of_lt_of_le {a b c d : Int} (h₁ : a < b) (h₂ : c ≤ d) : + a + c < b + d := + Int.lt_of_lt_of_le (Int.add_lt_add_right h₁ c) (Int.add_le_add_left h₂ b) + +protected theorem lt_add_of_pos_right (a : Int) {b : Int} (h : 0 < b) : a < a + b := by + have : a + 0 < a + b := Int.add_lt_add_left h a + rwa [Int.add_zero] at this + +protected theorem lt_add_of_pos_left (a : Int) {b : Int} (h : 0 < b) : a < b + a := by + have : 0 + a < b + a := Int.add_lt_add_right h a + rwa [Int.zero_add] at this + +protected theorem add_nonneg {a b : Int} (ha : 0 ≤ a) (hb : 0 ≤ b) : 0 ≤ a + b := + Int.zero_add 0 ▸ Int.add_le_add ha hb + +protected theorem add_pos {a b : Int} (ha : 0 < a) (hb : 0 < b) : 0 < a + b := + Int.zero_add 0 ▸ Int.add_lt_add ha hb + +protected theorem add_pos_of_pos_of_nonneg {a b : Int} (ha : 0 < a) (hb : 0 ≤ b) : 0 < a + b := + Int.zero_add 0 ▸ Int.add_lt_add_of_lt_of_le ha hb + +protected theorem add_pos_of_nonneg_of_pos {a b : Int} (ha : 0 ≤ a) (hb : 0 < b) : 0 < a + b := + Int.zero_add 0 ▸ Int.add_lt_add_of_le_of_lt ha hb + +protected theorem add_nonpos {a b : Int} (ha : a ≤ 0) (hb : b ≤ 0) : a + b ≤ 0 := + Int.zero_add 0 ▸ Int.add_le_add ha hb + +protected theorem add_neg {a b : Int} (ha : a < 0) (hb : b < 0) : a + b < 0 := + Int.zero_add 0 ▸ Int.add_lt_add ha hb + +protected theorem add_neg_of_neg_of_nonpos {a b : Int} (ha : a < 0) (hb : b ≤ 0) : a + b < 0 := + Int.zero_add 0 ▸ Int.add_lt_add_of_lt_of_le ha hb + +protected theorem add_neg_of_nonpos_of_neg {a b : Int} (ha : a ≤ 0) (hb : b < 0) : a + b < 0 := + Int.zero_add 0 ▸ Int.add_lt_add_of_le_of_lt ha hb + +protected theorem lt_add_of_le_of_pos {a b c : Int} (hbc : b ≤ c) (ha : 0 < a) : b < c + a := + Int.add_zero b ▸ Int.add_lt_add_of_le_of_lt hbc ha + +theorem add_one_le_iff {a b : Int} : a + 1 ≤ b ↔ a < b := .rfl + +theorem lt_add_one_iff {a b : Int} : a < b + 1 ↔ a ≤ b := Int.add_le_add_iff_right _ + +@[simp] theorem succ_ofNat_pos (n : Nat) : 0 < (n : Int) + 1 := + lt_add_one_iff.2 (ofNat_zero_le _) + +theorem le_add_one {a b : Int} (h : a ≤ b) : a ≤ b + 1 := + Int.le_of_lt (Int.lt_add_one_iff.2 h) + +protected theorem nonneg_of_neg_nonpos {a : Int} (h : -a ≤ 0) : 0 ≤ a := + Int.le_of_neg_le_neg <| by rwa [Int.neg_zero] + +protected theorem nonpos_of_neg_nonneg {a : Int} (h : 0 ≤ -a) : a ≤ 0 := + Int.le_of_neg_le_neg <| by rwa [Int.neg_zero] + +protected theorem lt_of_neg_lt_neg {a b : Int} (h : -b < -a) : a < b := + Int.neg_neg a ▸ Int.neg_neg b ▸ Int.neg_lt_neg h + +protected theorem pos_of_neg_neg {a : Int} (h : -a < 0) : 0 < a := + Int.lt_of_neg_lt_neg <| by rwa [Int.neg_zero] + +protected theorem neg_of_neg_pos {a : Int} (h : 0 < -a) : a < 0 := + have : -0 < -a := by rwa [Int.neg_zero] + Int.lt_of_neg_lt_neg this + +protected theorem le_neg_of_le_neg {a b : Int} (h : a ≤ -b) : b ≤ -a := by + have h := Int.neg_le_neg h + rwa [Int.neg_neg] at h + +protected theorem neg_le_of_neg_le {a b : Int} (h : -a ≤ b) : -b ≤ a := by + have h := Int.neg_le_neg h + rwa [Int.neg_neg] at h + +protected theorem lt_neg_of_lt_neg {a b : Int} (h : a < -b) : b < -a := by + have h := Int.neg_lt_neg h + rwa [Int.neg_neg] at h + +protected theorem neg_lt_of_neg_lt {a b : Int} (h : -a < b) : -b < a := by + have h := Int.neg_lt_neg h + rwa [Int.neg_neg] at h + +protected theorem sub_nonpos_of_le {a b : Int} (h : a ≤ b) : a - b ≤ 0 := by + have h := Int.add_le_add_right h (-b) + rwa [Int.add_right_neg] at h + +protected theorem le_of_sub_nonpos {a b : Int} (h : a - b ≤ 0) : a ≤ b := by + have h := Int.add_le_add_right h b + rwa [Int.sub_add_cancel, Int.zero_add] at h + +protected theorem sub_neg_of_lt {a b : Int} (h : a < b) : a - b < 0 := by + have h := Int.add_lt_add_right h (-b) + rwa [Int.add_right_neg] at h + +protected theorem lt_of_sub_neg {a b : Int} (h : a - b < 0) : a < b := by + have h := Int.add_lt_add_right h b + rwa [Int.sub_add_cancel, Int.zero_add] at h + +protected theorem add_le_of_le_neg_add {a b c : Int} (h : b ≤ -a + c) : a + b ≤ c := by + have h := Int.add_le_add_left h a + rwa [Int.add_neg_cancel_left] at h + +protected theorem le_neg_add_of_add_le {a b c : Int} (h : a + b ≤ c) : b ≤ -a + c := by + have h := Int.add_le_add_left h (-a) + rwa [Int.neg_add_cancel_left] at h + +protected theorem add_le_of_le_sub_left {a b c : Int} (h : b ≤ c - a) : a + b ≤ c := by + have h := Int.add_le_add_left h a + rwa [← Int.add_sub_assoc, Int.add_comm a c, Int.add_sub_cancel] at h + +protected theorem le_sub_left_of_add_le {a b c : Int} (h : a + b ≤ c) : b ≤ c - a := by + have h := Int.add_le_add_right h (-a) + rwa [Int.add_comm a b, Int.add_neg_cancel_right] at h + +protected theorem add_le_of_le_sub_right {a b c : Int} (h : a ≤ c - b) : a + b ≤ c := by + have h := Int.add_le_add_right h b + rwa [Int.sub_add_cancel] at h + +protected theorem le_sub_right_of_add_le {a b c : Int} (h : a + b ≤ c) : a ≤ c - b := by + have h := Int.add_le_add_right h (-b) + rwa [Int.add_neg_cancel_right] at h + +protected theorem le_add_of_neg_add_le {a b c : Int} (h : -b + a ≤ c) : a ≤ b + c := by + have h := Int.add_le_add_left h b + rwa [Int.add_neg_cancel_left] at h + +protected theorem neg_add_le_of_le_add {a b c : Int} (h : a ≤ b + c) : -b + a ≤ c := by + have h := Int.add_le_add_left h (-b) + rwa [Int.neg_add_cancel_left] at h + +protected theorem le_add_of_sub_left_le {a b c : Int} (h : a - b ≤ c) : a ≤ b + c := by + have h := Int.add_le_add_right h b + rwa [Int.sub_add_cancel, Int.add_comm] at h + +protected theorem le_add_of_sub_right_le {a b c : Int} (h : a - c ≤ b) : a ≤ b + c := by + have h := Int.add_le_add_right h c + rwa [Int.sub_add_cancel] at h + +protected theorem sub_right_le_of_le_add {a b c : Int} (h : a ≤ b + c) : a - c ≤ b := by + have h := Int.add_le_add_right h (-c) + rwa [Int.add_neg_cancel_right] at h + +protected theorem le_add_of_neg_add_le_left {a b c : Int} (h : -b + a ≤ c) : a ≤ b + c := by + rw [Int.add_comm] at h + exact Int.le_add_of_sub_left_le h + +protected theorem neg_add_le_left_of_le_add {a b c : Int} (h : a ≤ b + c) : -b + a ≤ c := by + rw [Int.add_comm] + exact Int.sub_left_le_of_le_add h + +protected theorem le_add_of_neg_add_le_right {a b c : Int} (h : -c + a ≤ b) : a ≤ b + c := by + rw [Int.add_comm] at h + exact Int.le_add_of_sub_right_le h + +protected theorem neg_add_le_right_of_le_add {a b c : Int} (h : a ≤ b + c) : -c + a ≤ b := by + rw [Int.add_comm] at h + exact Int.neg_add_le_left_of_le_add h + +protected theorem le_add_of_neg_le_sub_left {a b c : Int} (h : -a ≤ b - c) : c ≤ a + b := + Int.le_add_of_neg_add_le_left (Int.add_le_of_le_sub_right h) + +protected theorem neg_le_sub_left_of_le_add {a b c : Int} (h : c ≤ a + b) : -a ≤ b - c := by + have h := Int.le_neg_add_of_add_le (Int.sub_left_le_of_le_add h) + rwa [Int.add_comm] at h + +protected theorem le_add_of_neg_le_sub_right {a b c : Int} (h : -b ≤ a - c) : c ≤ a + b := + Int.le_add_of_sub_right_le (Int.add_le_of_le_sub_left h) + +protected theorem neg_le_sub_right_of_le_add {a b c : Int} (h : c ≤ a + b) : -b ≤ a - c := + Int.le_sub_left_of_add_le (Int.sub_right_le_of_le_add h) + +protected theorem sub_le_of_sub_le {a b c : Int} (h : a - b ≤ c) : a - c ≤ b := + Int.sub_left_le_of_le_add (Int.le_add_of_sub_right_le h) + +protected theorem sub_le_sub_left {a b : Int} (h : a ≤ b) (c : Int) : c - b ≤ c - a := + Int.add_le_add_left (Int.neg_le_neg h) c + +protected theorem sub_le_sub_right {a b : Int} (h : a ≤ b) (c : Int) : a - c ≤ b - c := + Int.add_le_add_right h (-c) + +protected theorem sub_le_sub {a b c d : Int} (hab : a ≤ b) (hcd : c ≤ d) : a - d ≤ b - c := + Int.add_le_add hab (Int.neg_le_neg hcd) + +protected theorem add_lt_of_lt_neg_add {a b c : Int} (h : b < -a + c) : a + b < c := by + have h := Int.add_lt_add_left h a + rwa [Int.add_neg_cancel_left] at h + +protected theorem lt_neg_add_of_add_lt {a b c : Int} (h : a + b < c) : b < -a + c := by + have h := Int.add_lt_add_left h (-a) + rwa [Int.neg_add_cancel_left] at h + +protected theorem add_lt_of_lt_sub_left {a b c : Int} (h : b < c - a) : a + b < c := by + have h := Int.add_lt_add_left h a + rwa [← Int.add_sub_assoc, Int.add_comm a c, Int.add_sub_cancel] at h + +protected theorem lt_sub_left_of_add_lt {a b c : Int} (h : a + b < c) : b < c - a := by + have h := Int.add_lt_add_right h (-a) + rwa [Int.add_comm a b, Int.add_neg_cancel_right] at h + +protected theorem add_lt_of_lt_sub_right {a b c : Int} (h : a < c - b) : a + b < c := by + have h := Int.add_lt_add_right h b + rwa [Int.sub_add_cancel] at h + +protected theorem lt_sub_right_of_add_lt {a b c : Int} (h : a + b < c) : a < c - b := by + have h := Int.add_lt_add_right h (-b) + rwa [Int.add_neg_cancel_right] at h + +protected theorem lt_add_of_neg_add_lt {a b c : Int} (h : -b + a < c) : a < b + c := by + have h := Int.add_lt_add_left h b + rwa [Int.add_neg_cancel_left] at h + +protected theorem neg_add_lt_of_lt_add {a b c : Int} (h : a < b + c) : -b + a < c := by + have h := Int.add_lt_add_left h (-b) + rwa [Int.neg_add_cancel_left] at h + +protected theorem lt_add_of_sub_left_lt {a b c : Int} (h : a - b < c) : a < b + c := by + have h := Int.add_lt_add_right h b + rwa [Int.sub_add_cancel, Int.add_comm] at h + +protected theorem sub_left_lt_of_lt_add {a b c : Int} (h : a < b + c) : a - b < c := by + have h := Int.add_lt_add_right h (-b) + rwa [Int.add_comm b c, Int.add_neg_cancel_right] at h + +protected theorem lt_add_of_sub_right_lt {a b c : Int} (h : a - c < b) : a < b + c := by + have h := Int.add_lt_add_right h c + rwa [Int.sub_add_cancel] at h + +protected theorem sub_right_lt_of_lt_add {a b c : Int} (h : a < b + c) : a - c < b := by + have h := Int.add_lt_add_right h (-c) + rwa [Int.add_neg_cancel_right] at h + +protected theorem lt_add_of_neg_add_lt_left {a b c : Int} (h : -b + a < c) : a < b + c := by + rw [Int.add_comm] at h + exact Int.lt_add_of_sub_left_lt h + +protected theorem neg_add_lt_left_of_lt_add {a b c : Int} (h : a < b + c) : -b + a < c := by + rw [Int.add_comm] + exact Int.sub_left_lt_of_lt_add h + +protected theorem lt_add_of_neg_add_lt_right {a b c : Int} (h : -c + a < b) : a < b + c := by + rw [Int.add_comm] at h + exact Int.lt_add_of_sub_right_lt h + +protected theorem neg_add_lt_right_of_lt_add {a b c : Int} (h : a < b + c) : -c + a < b := by + rw [Int.add_comm] at h + exact Int.neg_add_lt_left_of_lt_add h + +protected theorem lt_add_of_neg_lt_sub_left {a b c : Int} (h : -a < b - c) : c < a + b := + Int.lt_add_of_neg_add_lt_left (Int.add_lt_of_lt_sub_right h) + +protected theorem neg_lt_sub_left_of_lt_add {a b c : Int} (h : c < a + b) : -a < b - c := by + have h := Int.lt_neg_add_of_add_lt (Int.sub_left_lt_of_lt_add h) + rwa [Int.add_comm] at h + +protected theorem lt_add_of_neg_lt_sub_right {a b c : Int} (h : -b < a - c) : c < a + b := + Int.lt_add_of_sub_right_lt (Int.add_lt_of_lt_sub_left h) + +protected theorem neg_lt_sub_right_of_lt_add {a b c : Int} (h : c < a + b) : -b < a - c := + Int.lt_sub_left_of_add_lt (Int.sub_right_lt_of_lt_add h) + +protected theorem sub_lt_of_sub_lt {a b c : Int} (h : a - b < c) : a - c < b := + Int.sub_left_lt_of_lt_add (Int.lt_add_of_sub_right_lt h) + +protected theorem sub_lt_sub_left {a b : Int} (h : a < b) (c : Int) : c - b < c - a := + Int.add_lt_add_left (Int.neg_lt_neg h) c + +protected theorem sub_lt_sub_right {a b : Int} (h : a < b) (c : Int) : a - c < b - c := + Int.add_lt_add_right h (-c) + +protected theorem sub_lt_sub {a b c d : Int} (hab : a < b) (hcd : c < d) : a - d < b - c := + Int.add_lt_add hab (Int.neg_lt_neg hcd) + +protected theorem sub_lt_sub_of_le_of_lt {a b c d : Int} + (hab : a ≤ b) (hcd : c < d) : a - d < b - c := + Int.add_lt_add_of_le_of_lt hab (Int.neg_lt_neg hcd) + +protected theorem sub_lt_sub_of_lt_of_le {a b c d : Int} + (hab : a < b) (hcd : c ≤ d) : a - d < b - c := + Int.add_lt_add_of_lt_of_le hab (Int.neg_le_neg hcd) + +protected theorem add_le_add_three {a b c d e f : Int} + (h₁ : a ≤ d) (h₂ : b ≤ e) (h₃ : c ≤ f) : a + b + c ≤ d + e + f := + Int.add_le_add (Int.add_le_add h₁ h₂) h₃ + +theorem exists_eq_neg_ofNat {a : Int} (H : a ≤ 0) : ∃ n : Nat, a = -(n : Int) := + let ⟨n, h⟩ := eq_ofNat_of_zero_le (Int.neg_nonneg_of_nonpos H) + ⟨n, Int.eq_neg_of_eq_neg h.symm⟩ + +theorem lt_of_add_one_le {a b : Int} (H : a + 1 ≤ b) : a < b := H + +theorem lt_add_one_of_le {a b : Int} (H : a ≤ b) : a < b + 1 := Int.add_le_add_right H 1 + +theorem le_of_lt_add_one {a b : Int} (H : a < b + 1) : a ≤ b := Int.le_of_add_le_add_right H + +theorem sub_one_lt_of_le {a b : Int} (H : a ≤ b) : a - 1 < b := + Int.sub_right_lt_of_lt_add <| lt_add_one_of_le H + +theorem le_of_sub_one_lt {a b : Int} (H : a - 1 < b) : a ≤ b := + le_of_lt_add_one <| Int.lt_add_of_sub_right_lt H + +theorem le_sub_one_of_lt {a b : Int} (H : a < b) : a ≤ b - 1 := Int.le_sub_right_of_add_le H + +theorem lt_of_le_sub_one {a b : Int} (H : a ≤ b - 1) : a < b := Int.add_le_of_le_sub_right H + +/- ### Order properties and multiplication -/ + +protected theorem mul_lt_mul {a b c d : Int} + (h₁ : a < c) (h₂ : b ≤ d) (h₃ : 0 < b) (h₄ : 0 ≤ c) : a * b < c * d := + Int.lt_of_lt_of_le (Int.mul_lt_mul_of_pos_right h₁ h₃) (Int.mul_le_mul_of_nonneg_left h₂ h₄) + +protected theorem mul_lt_mul' {a b c d : Int} + (h₁ : a ≤ c) (h₂ : b < d) (h₃ : 0 ≤ b) (h₄ : 0 < c) : a * b < c * d := + Int.lt_of_le_of_lt (Int.mul_le_mul_of_nonneg_right h₁ h₃) (Int.mul_lt_mul_of_pos_left h₂ h₄) + +protected theorem mul_neg_of_pos_of_neg {a b : Int} (ha : 0 < a) (hb : b < 0) : a * b < 0 := by + have h : a * b < a * 0 := Int.mul_lt_mul_of_pos_left hb ha + rwa [Int.mul_zero] at h + +protected theorem mul_neg_of_neg_of_pos {a b : Int} (ha : a < 0) (hb : 0 < b) : a * b < 0 := by + have h : a * b < 0 * b := Int.mul_lt_mul_of_pos_right ha hb + rwa [Int.zero_mul] at h + +protected theorem mul_nonneg_of_nonpos_of_nonpos {a b : Int} + (ha : a ≤ 0) (hb : b ≤ 0) : 0 ≤ a * b := by + have : 0 * b ≤ a * b := Int.mul_le_mul_of_nonpos_right ha hb + rwa [Int.zero_mul] at this + +protected theorem mul_lt_mul_of_neg_left {a b c : Int} (h : b < a) (hc : c < 0) : c * a < c * b := + have : -c > 0 := Int.neg_pos_of_neg hc + have : -c * b < -c * a := Int.mul_lt_mul_of_pos_left h this + have : -(c * b) < -(c * a) := by + rwa [← Int.neg_mul_eq_neg_mul, ← Int.neg_mul_eq_neg_mul] at this + Int.lt_of_neg_lt_neg this + +protected theorem mul_lt_mul_of_neg_right {a b c : Int} (h : b < a) (hc : c < 0) : a * c < b * c := + have : -c > 0 := Int.neg_pos_of_neg hc + have : b * -c < a * -c := Int.mul_lt_mul_of_pos_right h this + have : -(b * c) < -(a * c) := by + rwa [← Int.neg_mul_eq_mul_neg, ← Int.neg_mul_eq_mul_neg] at this + Int.lt_of_neg_lt_neg this + +protected theorem mul_pos_of_neg_of_neg {a b : Int} (ha : a < 0) (hb : b < 0) : 0 < a * b := by + have : 0 * b < a * b := Int.mul_lt_mul_of_neg_right ha hb + rwa [Int.zero_mul] at this + +protected theorem mul_self_le_mul_self {a b : Int} (h1 : 0 ≤ a) (h2 : a ≤ b) : a * a ≤ b * b := + Int.mul_le_mul h2 h2 h1 (Int.le_trans h1 h2) + +protected theorem mul_self_lt_mul_self {a b : Int} (h1 : 0 ≤ a) (h2 : a < b) : a * a < b * b := + Int.mul_lt_mul' (Int.le_of_lt h2) h2 h1 (Int.lt_of_le_of_lt h1 h2) + +/- ## sign -/ + +@[simp] theorem sign_zero : sign 0 = 0 := rfl +@[simp] theorem sign_one : sign 1 = 1 := rfl +theorem sign_neg_one : sign (-1) = -1 := rfl + +@[simp] theorem sign_of_add_one (x : Nat) : Int.sign (x + 1) = 1 := rfl +@[simp] theorem sign_negSucc (x : Nat) : Int.sign (Int.negSucc x) = -1 := rfl + +theorem natAbs_sign (z : Int) : z.sign.natAbs = if z = 0 then 0 else 1 := + match z with | 0 | succ _ | -[_+1] => rfl + +theorem natAbs_sign_of_nonzero {z : Int} (hz : z ≠ 0) : z.sign.natAbs = 1 := by + rw [Int.natAbs_sign, if_neg hz] + +theorem sign_ofNat_of_nonzero {n : Nat} (hn : n ≠ 0) : Int.sign n = 1 := + match n, Nat.exists_eq_succ_of_ne_zero hn with + | _, ⟨n, rfl⟩ => Int.sign_of_add_one n + +@[simp] theorem sign_neg (z : Int) : Int.sign (-z) = -Int.sign z := by + match z with | 0 | succ _ | -[_+1] => rfl + +theorem sign_mul_natAbs : ∀ a : Int, sign a * natAbs a = a + | 0 => rfl + | succ _ => Int.one_mul _ + | -[_+1] => (Int.neg_eq_neg_one_mul _).symm + +@[simp] theorem sign_mul : ∀ a b, sign (a * b) = sign a * sign b + | a, 0 | 0, b => by simp [Int.mul_zero, Int.zero_mul] + | succ _, succ _ | succ _, -[_+1] | -[_+1], succ _ | -[_+1], -[_+1] => rfl + +theorem sign_eq_one_of_pos {a : Int} (h : 0 < a) : sign a = 1 := + match a, eq_succ_of_zero_lt h with + | _, ⟨_, rfl⟩ => rfl + +theorem sign_eq_neg_one_of_neg {a : Int} (h : a < 0) : sign a = -1 := + match a, eq_negSucc_of_lt_zero h with + | _, ⟨_, rfl⟩ => rfl + +theorem eq_zero_of_sign_eq_zero : ∀ {a : Int}, sign a = 0 → a = 0 + | 0, _ => rfl + +theorem pos_of_sign_eq_one : ∀ {a : Int}, sign a = 1 → 0 < a + | (_ + 1 : Nat), _ => ofNat_lt.2 (Nat.succ_pos _) + +theorem neg_of_sign_eq_neg_one : ∀ {a : Int}, sign a = -1 → a < 0 + | (_ + 1 : Nat), h => nomatch h + | 0, h => nomatch h + | -[_+1], _ => negSucc_lt_zero _ + +theorem sign_eq_one_iff_pos (a : Int) : sign a = 1 ↔ 0 < a := + ⟨pos_of_sign_eq_one, sign_eq_one_of_pos⟩ + +theorem sign_eq_neg_one_iff_neg (a : Int) : sign a = -1 ↔ a < 0 := + ⟨neg_of_sign_eq_neg_one, sign_eq_neg_one_of_neg⟩ + +@[simp] theorem sign_eq_zero_iff_zero (a : Int) : sign a = 0 ↔ a = 0 := + ⟨eq_zero_of_sign_eq_zero, fun h => by rw [h, sign_zero]⟩ + +@[simp] theorem sign_sign : sign (sign x) = sign x := by + match x with + | 0 => rfl + | .ofNat (_ + 1) => rfl + | .negSucc _ => rfl + +@[simp] theorem sign_nonneg : 0 ≤ sign x ↔ 0 ≤ x := by + match x with + | 0 => rfl + | .ofNat (_ + 1) => + simp (config := { decide := true }) only [sign, true_iff] + exact Int.le_add_one (ofNat_nonneg _) + | .negSucc _ => simp (config := { decide := true }) [sign] + +theorem mul_sign : ∀ i : Int, i * sign i = natAbs i + | succ _ => Int.mul_one _ + | 0 => Int.mul_zero _ + | -[_+1] => Int.mul_neg_one _ + +/- ## natAbs -/ + +theorem natAbs_ne_zero {a : Int} : a.natAbs ≠ 0 ↔ a ≠ 0 := not_congr Int.natAbs_eq_zero + +theorem natAbs_mul_self : ∀ {a : Int}, ↑(natAbs a * natAbs a) = a * a + | ofNat _ => rfl + | -[_+1] => rfl + +theorem eq_nat_or_neg (a : Int) : ∃ n : Nat, a = n ∨ a = -↑n := ⟨_, natAbs_eq a⟩ + +theorem natAbs_mul_natAbs_eq {a b : Int} {c : Nat} + (h : a * b = (c : Int)) : a.natAbs * b.natAbs = c := by rw [← natAbs_mul, h, natAbs] + +@[simp] theorem natAbs_mul_self' (a : Int) : (natAbs a * natAbs a : Int) = a * a := by + rw [← Int.ofNat_mul, natAbs_mul_self] + +theorem natAbs_eq_iff {a : Int} {n : Nat} : a.natAbs = n ↔ a = n ∨ a = -↑n := by + rw [← Int.natAbs_eq_natAbs_iff, Int.natAbs_ofNat] + +theorem natAbs_add_le (a b : Int) : natAbs (a + b) ≤ natAbs a + natAbs b := by + suffices ∀ a b : Nat, natAbs (subNatNat a b.succ) ≤ (a + b).succ by + match a, b with + | (a:Nat), (b:Nat) => rw [ofNat_add_ofNat, natAbs_ofNat]; apply Nat.le_refl + | (a:Nat), -[b+1] => rw [natAbs_ofNat, natAbs_negSucc]; apply this + | -[a+1], (b:Nat) => + rw [natAbs_negSucc, natAbs_ofNat, Nat.succ_add, Nat.add_comm a b]; apply this + | -[a+1], -[b+1] => rw [natAbs_negSucc, succ_add]; apply Nat.le_refl + refine fun a b => subNatNat_elim a b.succ + (fun m n i => n = b.succ → natAbs i ≤ (m + b).succ) ?_ + (fun i n (e : (n + i).succ = _) => ?_) rfl + · rintro i n rfl + rw [Nat.add_comm _ i, Nat.add_assoc] + exact Nat.le_add_right i (b.succ + b).succ + · apply succ_le_succ + rw [← succ.inj e, ← Nat.add_assoc, Nat.add_comm] + apply Nat.le_add_right + +theorem natAbs_sub_le (a b : Int) : natAbs (a - b) ≤ natAbs a + natAbs b := by + rw [← Int.natAbs_neg b]; apply natAbs_add_le + +theorem negSucc_eq' (m : Nat) : -[m+1] = -m - 1 := by simp only [negSucc_eq, Int.neg_add]; rfl + +theorem natAbs_lt_natAbs_of_nonneg_of_lt {a b : Int} + (w₁ : 0 ≤ a) (w₂ : a < b) : a.natAbs < b.natAbs := + match a, b, eq_ofNat_of_zero_le w₁, eq_ofNat_of_zero_le (Int.le_trans w₁ (Int.le_of_lt w₂)) with + | _, _, ⟨_, rfl⟩, ⟨_, rfl⟩ => ofNat_lt.1 w₂ + +theorem eq_natAbs_iff_mul_eq_zero : natAbs a = n ↔ (a - n) * (a + n) = 0 := by + rw [natAbs_eq_iff, Int.mul_eq_zero, ← Int.sub_neg, Int.sub_eq_zero, Int.sub_eq_zero] + +end Int diff --git a/src/Init/Data/Int/Pow.lean b/src/Init/Data/Int/Pow.lean new file mode 100644 index 000000000000..f19f418aac4c --- /dev/null +++ b/src/Init/Data/Int/Pow.lean @@ -0,0 +1,44 @@ +/- +Copyright (c) 2016 Jeremy Avigad. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Avigad, Deniz Aydin, Floris van Doorn, Mario Carneiro +-/ +prelude +import Init.Data.Int.Lemmas + +namespace Int + +/-! # pow -/ + +protected theorem pow_zero (b : Int) : b^0 = 1 := rfl + +protected theorem pow_succ (b : Int) (e : Nat) : b ^ (e+1) = (b ^ e) * b := rfl +protected theorem pow_succ' (b : Int) (e : Nat) : b ^ (e+1) = b * (b ^ e) := by + rw [Int.mul_comm, Int.pow_succ] + +theorem pow_le_pow_of_le_left {n m : Nat} (h : n ≤ m) : ∀ (i : Nat), n^i ≤ m^i + | 0 => Nat.le_refl _ + | i + 1 => Nat.mul_le_mul (pow_le_pow_of_le_left h i) h + +theorem pow_le_pow_of_le_right {n : Nat} (hx : n > 0) {i : Nat} : ∀ {j}, i ≤ j → n^i ≤ n^j + | 0, h => + have : i = 0 := Nat.eq_zero_of_le_zero h + this.symm ▸ Nat.le_refl _ + | j + 1, h => + match Nat.le_or_eq_of_le_succ h with + | Or.inl h => show n^i ≤ n^j * n from + have : n^i * 1 ≤ n^j * n := Nat.mul_le_mul (pow_le_pow_of_le_right hx h) hx + Nat.mul_one (n^i) ▸ this + | Or.inr h => + h.symm ▸ Nat.le_refl _ + +theorem pos_pow_of_pos {n : Nat} (m : Nat) (h : 0 < n) : 0 < n^m := + pow_le_pow_of_le_right h (Nat.zero_le _) + +theorem natCast_pow (b n : Nat) : ((b^n : Nat) : Int) = (b : Int) ^ n := by + match n with + | 0 => rfl + | n + 1 => + simp only [Nat.pow_succ, Int.pow_succ, natCast_mul, natCast_pow _ n] + +end Int diff --git a/src/Init/Data/Nat/Basic.lean b/src/Init/Data/Nat/Basic.lean index 49d4f35f9f17..8bee45a010ef 100644 --- a/src/Init/Data/Nat/Basic.lean +++ b/src/Init/Data/Nat/Basic.lean @@ -232,6 +232,9 @@ protected theorem mul_assoc : ∀ (n m k : Nat), (n * m) * k = n * (m * k) protected theorem mul_left_comm (n m k : Nat) : n * (m * k) = m * (n * k) := by rw [← Nat.mul_assoc, Nat.mul_comm n m, Nat.mul_assoc] +protected theorem mul_two (n) : n * 2 = n + n := by rw [Nat.mul_succ, Nat.mul_one] +protected theorem two_mul (n) : 2 * n = n + n := by rw [Nat.succ_mul, Nat.one_mul] + /-! # Inequalities -/ attribute [simp] Nat.le_refl diff --git a/src/Init/Data/Nat/Bitwise/Lemmas.lean b/src/Init/Data/Nat/Bitwise/Lemmas.lean index ba554b586781..e8593f1355d7 100644 --- a/src/Init/Data/Nat/Bitwise/Lemmas.lean +++ b/src/Init/Data/Nat/Bitwise/Lemmas.lean @@ -6,6 +6,7 @@ Authors: Joe Hendrix prelude import Init.Data.Bool +import Init.Data.Int.Pow import Init.Data.Nat.Bitwise.Basic import Init.Data.Nat.Lemmas import Init.TacticsExtra diff --git a/src/Init/Data/Nat/Div.lean b/src/Init/Data/Nat/Div.lean index e51da62f2c6a..eb6be7c375cd 100644 --- a/src/Init/Data/Nat/Div.lean +++ b/src/Init/Data/Nat/Div.lean @@ -10,6 +10,13 @@ import Init.Data.Nat.Basic namespace Nat +/-- +Divisibility of natural numbers. `a ∣ b` (typed as `\|`) says that +there is some `c` such that `b = a * c`. +-/ +instance : Dvd Nat where + dvd a b := Exists (fun c => b = a * c) + theorem div_rec_lemma {x y : Nat} : 0 < y ∧ y ≤ x → x - y < x := fun ⟨ypos, ylex⟩ => sub_lt (Nat.lt_of_lt_of_le ypos ylex) ypos diff --git a/src/Init/Data/Nat/Dvd.lean b/src/Init/Data/Nat/Dvd.lean index ad3eff53089e..6732819ccc2a 100644 --- a/src/Init/Data/Nat/Dvd.lean +++ b/src/Init/Data/Nat/Dvd.lean @@ -5,17 +5,10 @@ Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ prelude import Init.Data.Nat.Div -import Init.TacticsExtra +import Init.Meta namespace Nat -/-- -Divisibility of natural numbers. `a ∣ b` (typed as `\|`) says that -there is some `c` such that `b = a * c`. --/ -instance : Dvd Nat where - dvd a b := Exists (fun c => b = a * c) - protected theorem dvd_refl (a : Nat) : a ∣ a := ⟨1, by simp⟩ protected theorem dvd_zero (a : Nat) : a ∣ 0 := ⟨0, by simp⟩ diff --git a/src/Init/Data/Nat/Lemmas.lean b/src/Init/Data/Nat/Lemmas.lean index 8ffa17d876d2..cdc3e696f9f6 100644 --- a/src/Init/Data/Nat/Lemmas.lean +++ b/src/Init/Data/Nat/Lemmas.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ prelude -import Init.Data.Nat.Dvd import Init.Data.Nat.MinMax import Init.Data.Nat.Log2 import Init.Data.Nat.Power2 @@ -410,10 +409,6 @@ protected theorem mul_right_comm (n m k : Nat) : n * m * k = n * k * m := by protected theorem mul_mul_mul_comm (a b c d : Nat) : (a * b) * (c * d) = (a * c) * (b * d) := by rw [Nat.mul_assoc, Nat.mul_assoc, Nat.mul_left_comm b] -protected theorem mul_two (n) : n * 2 = n + n := by rw [Nat.mul_succ, Nat.mul_one] - -protected theorem two_mul (n) : 2 * n = n + n := by rw [Nat.succ_mul, Nat.one_mul] - theorem mul_eq_zero : ∀ {m n}, n * m = 0 ↔ n = 0 ∨ m = 0 | 0, _ => ⟨fun _ => .inr rfl, fun _ => rfl⟩ | _, 0 => ⟨fun _ => .inl rfl, fun _ => Nat.zero_mul ..⟩ diff --git a/src/Init/Data/Random.lean b/src/Init/Data/Random.lean index 6c798d2a414f..b5475ba9d7f9 100644 --- a/src/Init/Data/Random.lean +++ b/src/Init/Data/Random.lean @@ -5,7 +5,6 @@ Authors: Leonardo de Moura -/ prelude import Init.System.IO -import Init.Data.Int universe u /-! diff --git a/src/Init/Omega/Int.lean b/src/Init/Omega/Int.lean index a01507212f01..f5765a053f3c 100644 --- a/src/Init/Omega/Int.lean +++ b/src/Init/Omega/Int.lean @@ -6,7 +6,6 @@ Authors: Scott Morrison prelude import Init.Data.Int.DivMod import Init.Data.Int.Order -import Init.Data.Nat.Basic /-! # Lemmas about `Nat`, `Int`, and `Fin` needed internally by `omega`. diff --git a/src/Init/Omega/IntList.lean b/src/Init/Omega/IntList.lean index 1d01cd38b1f0..3813fe5c30b2 100644 --- a/src/Init/Omega/IntList.lean +++ b/src/Init/Omega/IntList.lean @@ -6,7 +6,7 @@ Authors: Scott Morrison prelude import Init.Data.List.Lemmas import Init.Data.Int.DivModLemmas -import Init.Data.Int.Gcd +import Init.Data.Nat.Gcd namespace Lean.Omega diff --git a/src/Lean/Elab/Tactic/Omega/MinNatAbs.lean b/src/Lean/Elab/Tactic/Omega/MinNatAbs.lean index 0acd798c567e..3a227440c84e 100644 --- a/src/Lean/Elab/Tactic/Omega/MinNatAbs.lean +++ b/src/Lean/Elab/Tactic/Omega/MinNatAbs.lean @@ -5,8 +5,10 @@ Authors: Scott Morrison -/ prelude import Init.BinderPredicates +import Init.Data.Int.Order +import Init.Data.List.Lemmas +import Init.Data.Nat.MinMax import Init.Data.Option.Lemmas -import Init.Data.Nat.Bitwise.Lemmas /-! # `List.nonzeroMinimum`, `List.minNatAbs`, `List.maxNatAbs` diff --git a/src/Lean/Elab/Tactic/Omega/OmegaM.lean b/src/Lean/Elab/Tactic/Omega/OmegaM.lean index 6e58833f1372..8f0a28c00531 100644 --- a/src/Lean/Elab/Tactic/Omega/OmegaM.lean +++ b/src/Lean/Elab/Tactic/Omega/OmegaM.lean @@ -7,7 +7,7 @@ prelude import Init.Omega.LinearCombo import Init.Omega.Int import Init.Omega.Logic -import Init.Data.BitVec +import Init.Data.BitVec.Basic import Lean.Meta.AppBuilder /-! @@ -176,7 +176,7 @@ def analyzeAtom (e : Expr) : OmegaM (HashSet Expr) := do | _, (``Fin.val, #[n, i]) => r := r.insert (mkApp2 (.const ``Fin.isLt []) n i) | _, (``BitVec.toNat, #[n, x]) => - r := r.insert (mkApp2 (.const ``BitVec.toNat_lt []) n x) + r := r.insert (mkApp2 (.const ``BitVec.isLt []) n x) | _, _ => pure () return r | (``HDiv.hDiv, #[_, _, _, _, x, k]) => match natCast? k with From b2ae4bd5c1a4b3b3cfe67d72be9397459147d8c7 Mon Sep 17 00:00:00 2001 From: Mac Malone Date: Mon, 11 Mar 2024 22:46:42 -0400 Subject: [PATCH 15/32] feat: allow `noncomputable unsafe` definitions (#3647) Enables the combination of `noncomputable unsafe` to be used for definitions. Outside of pure theory, `noncomputable` is also useful to prevent Lean from compiling a definition which will be implemented with external code later. Such definitions may also wish to be marked `unsafe` if they perform morally impure or memory-unsafe functions. --- src/Lean/Elab/MutualDef.lean | 2 -- 1 file changed, 2 deletions(-) diff --git a/src/Lean/Elab/MutualDef.lean b/src/Lean/Elab/MutualDef.lean index edb8e7e94fc8..d58193aa5eb7 100644 --- a/src/Lean/Elab/MutualDef.lean +++ b/src/Lean/Elab/MutualDef.lean @@ -68,8 +68,6 @@ private def check (prevHeaders : Array DefViewElabHeader) (newHeader : DefViewEl throwError "'partial' theorems are not allowed, 'partial' is a code generation directive" if newHeader.kind.isTheorem && newHeader.modifiers.isNoncomputable then throwError "'theorem' subsumes 'noncomputable', code is not generated for theorems" - if newHeader.modifiers.isNoncomputable && newHeader.modifiers.isUnsafe then - throwError "'noncomputable unsafe' is not allowed" if newHeader.modifiers.isNoncomputable && newHeader.modifiers.isPartial then throwError "'noncomputable partial' is not allowed" if newHeader.modifiers.isPartial && newHeader.modifiers.isUnsafe then From 5cf4db7fbf1616a205b2f263689fd750c1ab8444 Mon Sep 17 00:00:00 2001 From: thorimur <68410468+thorimur@users.noreply.github.com> Date: Tue, 12 Mar 2024 01:17:58 -0400 Subject: [PATCH 16/32] fix: make `dsimp?` use and report simprocs (#3654) Modifies `dsimpLocation'` (which implements `dsimp?`) to take a `simprocs : SimprocsArray` argument, like `simpLocation` and `dsimpLocation`. This ensures that the behavior of `dsimp` matches `dsimp?`. --- Closes #3653 --- src/Lean/Elab/Tactic/SimpTrace.lean | 12 +++++++----- 1 file changed, 7 insertions(+), 5 deletions(-) diff --git a/src/Lean/Elab/Tactic/SimpTrace.lean b/src/Lean/Elab/Tactic/SimpTrace.lean index 7d7ed2518ddf..76d2a7ece49a 100644 --- a/src/Lean/Elab/Tactic/SimpTrace.lean +++ b/src/Lean/Elab/Tactic/SimpTrace.lean @@ -56,7 +56,8 @@ def mkSimpCallStx (stx : Syntax) (usedSimps : UsedSimps) : MetaM (TSyntax `tacti | _ => throwUnsupportedSyntax /-- Implementation of `dsimp?`. -/ -def dsimpLocation' (ctx : Simp.Context) (loc : Location) : TacticM Simp.UsedSimps := do +def dsimpLocation' (ctx : Simp.Context) (simprocs : SimprocsArray) (loc : Location) : + TacticM Simp.UsedSimps := do match loc with | Location.targets hyps simplifyTarget => withMainContext do @@ -69,8 +70,8 @@ where /-- Implementation of `dsimp?`. -/ go (fvarIdsToSimp : Array FVarId) (simplifyTarget : Bool) : TacticM Simp.UsedSimps := do let mvarId ← getMainGoal - let (result?, usedSimps) ← - dsimpGoal mvarId ctx (simplifyTarget := simplifyTarget) (fvarIdsToSimp := fvarIdsToSimp) + let (result?, usedSimps) ← dsimpGoal mvarId ctx simprocs (simplifyTarget := simplifyTarget) + (fvarIdsToSimp := fvarIdsToSimp) match result? with | none => replaceMainGoal [] | some mvarId => replaceMainGoal [mvarId] @@ -83,8 +84,9 @@ where `(tactic| dsimp!%$tk $(config)? $[only%$o]? $[[$args,*]]? $(loc)?) else `(tactic| dsimp%$tk $(config)? $[only%$o]? $[[$args,*]]? $(loc)?) - let { ctx, .. } ← withMainContext <| mkSimpContext stx (eraseLocal := false) (kind := .dsimp) - let usedSimps ← dsimpLocation' ctx <| (loc.map expandLocation).getD (.targets #[] true) + let { ctx, simprocs, .. } ← + withMainContext <| mkSimpContext stx (eraseLocal := false) (kind := .dsimp) + let usedSimps ← dsimpLocation' ctx simprocs <| (loc.map expandLocation).getD (.targets #[] true) let stx ← mkSimpCallStx stx usedSimps addSuggestion tk stx (origSpan? := ← getRef) | _ => throwUnsupportedSyntax From 07dac6784705fdd4a5e060ca569f247bed7b3f58 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Tue, 12 Mar 2024 17:35:14 +0100 Subject: [PATCH 17/32] feat: guard_msgs to escapes trailing newlines (#3617) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This makes trailing whitespace visible and protectes them against trimming by the editor, by appending the symbol ⏎ to such a line (and also to any line that ends with such a symbol, to avoid ambiguities in the case the message already had that symbol). (Only the code action output / docstring parsing is affected; the error message as sent to the InfoView is unaffected.) Fixes #3571 --- src/Lean/Elab/GuardMsgs.lean | 25 +++++++++++++++--- tests/lean/run/guard_msgs.lean | 46 ++++++++++++++++++++++++++++++++++ 2 files changed, 68 insertions(+), 3 deletions(-) diff --git a/src/Lean/Elab/GuardMsgs.lean b/src/Lean/Elab/GuardMsgs.lean index b56889bea9ac..ecdca448c4ce 100644 --- a/src/Lean/Elab/GuardMsgs.lean +++ b/src/Lean/Elab/GuardMsgs.lean @@ -73,9 +73,28 @@ structure GuardMsgFailure where res : String deriving TypeName +/-- +Makes trailing whitespace visible and protectes them against trimming by the editor, by appending +the symbol ⏎ to such a line (and also to any line that ends with such a symbol, to avoid +ambiguities in the case the message already had that symbol). +-/ +def revealTrailingWhitespace (s : String) : String := + s.replace "⏎\n" "⏎⏎\n" |>.replace "\t\n" "\t⏎\n" |>.replace " \n" " ⏎\n" + +/- The inverse of `revealTrailingWhitespace` -/ +def removeTrailingWhitespaceMarker (s : String) : String := + s.replace "⏎\n" "\n" + +/-- +Strings are compared up to newlines, to allow users to break long lines. +-/ +def equalUpToNewlines (exp res : String) : Bool := + exp.replace "\n" " " == res.replace "\n" " " + @[builtin_command_elab Lean.guardMsgsCmd] def elabGuardMsgs : CommandElab | `(command| $[$dc?:docComment]? #guard_msgs%$tk $(spec?)? in $cmd) => do - let expected : String := (← dc?.mapM (getDocStringText ·)).getD "" |>.trim + let expected : String := (← dc?.mapM (getDocStringText ·)).getD "" + |>.trim |> removeTrailingWhitespaceMarker let specFn ← parseGuardMsgsSpec spec? let initMsgs ← modifyGet fun st => (st.messages, { st with messages := {} }) elabCommandTopLevel cmd @@ -88,8 +107,7 @@ deriving TypeName | .drop => pure () | .passthrough => toPassthrough := toPassthrough.add msg let res := "---\n".intercalate (← toCheck.toList.mapM (messageToStringWithoutPos ·)) |>.trim - -- We do some whitespace normalization here to allow users to break long lines. - if expected.replace "\n" " " == res.replace "\n" " " then + if equalUpToNewlines expected res then -- Passed. Only put toPassthrough messages back on the message log modify fun st => { st with messages := initMsgs ++ toPassthrough } else @@ -119,6 +137,7 @@ def guardMsgsCodeAction : CommandCodeAction := fun _ _ _ node => do lazy? := some do let some start := stx.getPos? true | return eager let some tail := stx.setArg 0 mkNullNode |>.getPos? true | return eager + let res := revealTrailingWhitespace res let newText := if res.isEmpty then "" else if res.length ≤ 100-7 && !res.contains '\n' then -- TODO: configurable line length? diff --git a/tests/lean/run/guard_msgs.lean b/tests/lean/run/guard_msgs.lean index c7e85438e1fa..2e4faa99dc59 100644 --- a/tests/lean/run/guard_msgs.lean +++ b/tests/lean/run/guard_msgs.lean @@ -1,3 +1,5 @@ +import Lean.Elab.Command + #guard_msgs in /-- error: unknown identifier 'x' -/ #guard_msgs in @@ -49,3 +51,47 @@ error: failed to synthesize instance -/ #guard_msgs(error) in example : α := 22 + +-- Trailing whitespace + +/-- +info: foo ⏎ +bar +--- +error: ❌ Docstring on `#guard_msgs` does not match generated message: + +info: foo ⏎ +bar +-/ +#guard_msgs in +#guard_msgs in +run_meta Lean.logInfo "foo \nbar" + +#guard_msgs in +/-- +info: foo ⏎ +bar +-/ +#guard_msgs in +run_meta Lean.logInfo "foo \nbar" + +/-- +info: foo ⏎⏎ +bar +--- +error: ❌ Docstring on `#guard_msgs` does not match generated message: + +info: foo ⏎⏎ +bar +-/ +#guard_msgs in +#guard_msgs in +run_meta Lean.logInfo "foo ⏎\nbar" + +#guard_msgs in +/-- +info: foo ⏎⏎ +bar +-/ +#guard_msgs in +run_meta Lean.logInfo "foo ⏎\nbar" From 5aca09abca18206ac250a16b15e877d2f834fd28 Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Tue, 12 Mar 2024 16:18:56 -0700 Subject: [PATCH 18/32] fix: add `Canonicalizer.lean` and use it to canonicalize terms in `omega` (#3639) --- src/Lean/Elab/Tactic/Omega/OmegaM.lean | 6 +- src/Lean/Meta.lean | 1 + src/Lean/Meta/Canonicalizer.lean | 146 +++++++++++++++++++++++++ tests/lean/run/omegaCanon.lean | 17 +++ 4 files changed, 168 insertions(+), 2 deletions(-) create mode 100644 src/Lean/Meta/Canonicalizer.lean create mode 100644 tests/lean/run/omegaCanon.lean diff --git a/src/Lean/Elab/Tactic/Omega/OmegaM.lean b/src/Lean/Elab/Tactic/Omega/OmegaM.lean index 8f0a28c00531..57fd2bdcc4e2 100644 --- a/src/Lean/Elab/Tactic/Omega/OmegaM.lean +++ b/src/Lean/Elab/Tactic/Omega/OmegaM.lean @@ -9,6 +9,7 @@ import Init.Omega.Int import Init.Omega.Logic import Init.Data.BitVec.Basic import Lean.Meta.AppBuilder +import Lean.Meta.Canonicalizer /-! # The `OmegaM` state monad. @@ -54,7 +55,7 @@ structure State where atoms : HashMap Expr Nat := {} /-- An intermediate layer in the `OmegaM` monad. -/ -abbrev OmegaM' := StateRefT State (ReaderT Context MetaM) +abbrev OmegaM' := StateRefT State (ReaderT Context CanonM) /-- Cache of expressions that have been visited, and their reflection as a linear combination. @@ -70,7 +71,7 @@ abbrev OmegaM := StateRefT Cache OmegaM' /-- Run a computation in the `OmegaM` monad, starting with no recorded atoms. -/ def OmegaM.run (m : OmegaM α) (cfg : OmegaConfig) : MetaM α := - m.run' HashMap.empty |>.run' {} { cfg } + m.run' HashMap.empty |>.run' {} { cfg } |>.run /-- Retrieve the user-specified configuration options. -/ def cfg : OmegaM OmegaConfig := do pure (← read).cfg @@ -244,6 +245,7 @@ Return its index, and, if it is new, a collection of interesting facts about the -/ def lookup (e : Expr) : OmegaM (Nat × Option (HashSet Expr)) := do let c ← getThe State + let e ← canon e match c.atoms.find? e with | some i => return (i, none) | none => diff --git a/src/Lean/Meta.lean b/src/Lean/Meta.lean index 02c39e2de62e..32072eeff9c7 100644 --- a/src/Lean/Meta.lean +++ b/src/Lean/Meta.lean @@ -48,3 +48,4 @@ import Lean.Meta.Iterator import Lean.Meta.LazyDiscrTree import Lean.Meta.LitValues import Lean.Meta.CheckTactic +import Lean.Meta.Canonicalizer diff --git a/src/Lean/Meta/Canonicalizer.lean b/src/Lean/Meta/Canonicalizer.lean new file mode 100644 index 000000000000..ba98f1029aa6 --- /dev/null +++ b/src/Lean/Meta/Canonicalizer.lean @@ -0,0 +1,146 @@ +/- +Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Leonardo de Moura +-/ +prelude +import Lean.Util.ShareCommon +import Lean.Data.HashMap +import Lean.Meta.Basic +import Lean.Meta.FunInfo + +namespace Lean.Meta +namespace Canonicalizer + +/-! +Applications have implicit arguments. Thus, two terms that may look identical when pretty-printed can be structurally different. +For example, `@id (Id Nat) x` and `@id Nat x` are structurally different but are both pretty-printed as `id x`. +Moreover, these two terms are definitionally equal since `Id Nat` reduces to `Nat`. This may create situations +that are counterintuitive to our users. Furthermore, several tactics (e.g., `omega`) need to collect unique atoms in a goal. +One simple approach is to maintain a list of atoms found so far, and whenever a new atom is discovered, perform a +linear scan to test whether it is definitionally equal to a previously found one. However, this method is too costly, +even if the definitional equality test were inexpensive. + +This module aims to efficiently identify terms that are structurally different, definitionally equal, and structurally equal +when we disregard implicit arguments like `@id (Id Nat) x` and `@id Nat x`. The procedure is straightforward. For each atom, +we create a new abstracted atom by erasing all implicit information. We refer to this abstracted atom as a 'key.' For the two +terms mentioned, the key would be `@id _ x`, where `_` denotes a placeholder for a dummy term. To preserve any +pre-existing directed acyclic graph (DAG) structure and prevent exponential blowups while constructing the key, we employ +unsafe techniques, such as pointer equality. Additionally, we maintain a mapping from keys to lists of terms, where each +list contains terms sharing the same key but not definitionally equal. We posit that these lists will be small in practice. +-/ + +/-- +Auxiliary structure for creating a pointer-equality mapping from `Expr` to `Key`. +We use this mapping to ensure we preserve the dag-structure of input expressions. +-/ +structure ExprVisited where + e : Expr + deriving Inhabited + +unsafe instance : BEq ExprVisited where + beq a b := ptrAddrUnsafe a == ptrAddrUnsafe b + +unsafe instance : Hashable ExprVisited where + hash a := USize.toUInt64 (ptrAddrUnsafe a) + +abbrev Key := ExprVisited + +/-- +State for the `CanonM` monad. +-/ +structure State where + /-- "Set" of all keys created so far. This is a hash-consing helper structure available in Lean. -/ + keys : ShareCommon.State.{0} Lean.ShareCommon.objectFactory := ShareCommon.State.mk Lean.ShareCommon.objectFactory + /-- Mapping from `Expr` to `Key`. See comment at `ExprVisited`. -/ + -- We use `HashMapImp` to ensure we don't have to tag `State` as `unsafe`. + cache : HashMapImp ExprVisited Key := mkHashMapImp + /-- + Given a key `k` and `keyToExprs.find? k = some es`, we have that all `es` share key `k`, and + are not definitionally equal modulo the transparency setting used. -/ + keyToExprs : HashMapImp Key (List Expr) := mkHashMapImp + +instance : Inhabited State where + default := {} + +abbrev CanonM := ReaderT TransparencyMode $ StateRefT State MetaM + +/-- +The definitionally equality tests are performed using the given transparency mode. +We claim `TransparencyMode.instances` is a good setting for most applications. +-/ +def CanonM.run (x : CanonM α) (transparency := TransparencyMode.instances) : MetaM α := + StateRefT'.run' (x transparency) {} + +private def shareCommon (a : α) : CanonM α := + modifyGet fun { keys, cache, keyToExprs } => + let (a, keys) := ShareCommon.State.shareCommon keys a + (a, { keys, cache, keyToExprs }) + +private partial def mkKey (e : Expr) : CanonM Key := do + if let some key := unsafe (← get).cache.find? { e } then + return key + else + let key ← match e with + | .sort .. | .fvar .. | .bvar .. | .const .. | .lit .. => + pure { e := (← shareCommon e) } + | .mvar .. => + -- We instantiate assigned metavariables because the + -- pretty-printer also instantiates them. + let eNew ← instantiateMVars e + if eNew == e then pure { e := (← shareCommon e) } + else mkKey eNew + | .mdata _ a => mkKey a + | .app .. => + let f := (← mkKey e.getAppFn).e + if f.isMVar then + let eNew ← instantiateMVars e + unless eNew == e do + return (← mkKey eNew) + let info ← getFunInfo f + let args ← e.getAppArgs.mapIdxM fun i arg => do + if h : i < info.paramInfo.size then + let info := info.paramInfo[i] + if info.isExplicit then + pure (← mkKey arg).e + else + pure (mkSort 0) -- some dummy value for erasing implicit + else + pure (← mkKey arg).e + pure { e := (← shareCommon (mkAppN f args)) } + | .lam n t b i => + pure { e := (← shareCommon (.lam n (← mkKey t).e (← mkKey b).e i)) } + | .forallE n t b i => + pure { e := (← shareCommon (.forallE n (← mkKey t).e (← mkKey b).e i)) } + | .letE n t v b d => + pure { e := (← shareCommon (.letE n (← mkKey t).e (← mkKey v).e (← mkKey b).e d)) } + | .proj t i s => + pure { e := (← shareCommon (.proj t i (← mkKey s).e)) } + unsafe modify fun { keys, cache, keyToExprs} => { keys, keyToExprs, cache := cache.insert { e } key |>.1 } + return key + +/-- +"Canonicalize" the given expression. +-/ +def canon (e : Expr) : CanonM Expr := do + let k ← mkKey e + -- Find all expressions canonicalized before that have the same key. + if let some es' := unsafe (← get).keyToExprs.find? k then + withTransparency (← read) do + for e' in es' do + -- Found an expression `e'` that is definitionally equal to `e` and share the same key. + if (← isDefEq e e') then + return e' + -- `e` is not definitionally equal to any expression in `es'`. We claim this should be rare. + unsafe modify fun { keys, cache, keyToExprs } => { keys, cache, keyToExprs := keyToExprs.insert k (e :: es') |>.1 } + return e + else + -- `e` is the first expression we found with key `k`. + unsafe modify fun { keys, cache, keyToExprs } => { keys, cache, keyToExprs := keyToExprs.insert k [e] |>.1 } + return e + +end Canonicalizer + +export Canonicalizer (CanonM canon) + +end Lean.Meta diff --git a/tests/lean/run/omegaCanon.lean b/tests/lean/run/omegaCanon.lean new file mode 100644 index 000000000000..140b73113b87 --- /dev/null +++ b/tests/lean/run/omegaCanon.lean @@ -0,0 +1,17 @@ +def filter (p : α → Prop) [inst : DecidablePred p] (xs : List α) : List α := + match xs with + | [] => [] + | x :: xs' => + if p x then + -- Trying to confuse `omega` by creating subterms that are structurally different + -- but definitionally equal. + x :: @filter α p (fun x => inst x) xs' + else + filter p xs' + +def filter_length (p : α → Prop) [DecidablePred p] : (filter p xs).length ≤ xs.length := by + induction xs with + | nil => simp [filter] + | cons x xs ih => + simp only [filter] + split <;> simp only [List.length] <;> omega From 317adf42e92a4dbe07295bc1f0429b61bb8079fe Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 13 Mar 2024 16:35:52 +1100 Subject: [PATCH 19/32] chore: add @[simp] to Nat.succ_eq_add_one, and cleanup downstream (#3579) --- src/Init/Data/BitVec/Lemmas.lean | 2 +- src/Init/Data/Int/DivModLemmas.lean | 2 +- src/Init/Data/Int/Lemmas.lean | 25 +++++++++---------- src/Init/Data/Nat/Basic.lean | 18 ++++++++----- src/Init/Data/Nat/Bitwise/Basic.lean | 2 +- src/Init/Data/Nat/Div.lean | 2 +- src/Init/Data/Nat/Lemmas.lean | 3 +-- src/Init/Prelude.lean | 3 ++- tests/lean/guessLexFailures.lean.expected.out | 2 +- tests/lean/matchMultAlt.lean | 2 +- tests/lean/run/ac_rfl.lean | 2 +- tests/lean/run/eqnsAtSimp.lean | 2 +- tests/lean/run/eqnsAtSimp2.lean | 2 +- tests/lean/run/simp3.lean | 1 - tests/lean/run/simpRwBug.lean | 2 +- tests/lean/run/where1.lean | 2 +- 16 files changed, 38 insertions(+), 34 deletions(-) diff --git a/src/Init/Data/BitVec/Lemmas.lean b/src/Init/Data/BitVec/Lemmas.lean index 7ad648eb102f..3a965f1af8a0 100644 --- a/src/Init/Data/BitVec/Lemmas.lean +++ b/src/Init/Data/BitVec/Lemmas.lean @@ -70,7 +70,7 @@ theorem eq_of_getMsb_eq {x y : BitVec w} else have w_pos := Nat.pos_of_ne_zero w_zero have r : i ≤ w - 1 := by - simp [Nat.le_sub_iff_add_le w_pos, Nat.add_succ] + simp [Nat.le_sub_iff_add_le w_pos] exact i_lt have q_lt : w - 1 - i < w := by simp only [Nat.sub_sub] diff --git a/src/Init/Data/Int/DivModLemmas.lean b/src/Init/Data/Int/DivModLemmas.lean index 6ef222542037..19162a6a6138 100644 --- a/src/Init/Data/Int/DivModLemmas.lean +++ b/src/Init/Data/Int/DivModLemmas.lean @@ -760,7 +760,7 @@ theorem ediv_eq_ediv_of_mul_eq_mul {a b c d : Int} @[simp] protected theorem div_one : ∀ a : Int, a.div 1 = a | (n:Nat) => congrArg ofNat (Nat.div_one _) - | -[n+1] => by simp [Int.div, neg_ofNat_succ] + | -[n+1] => by simp [Int.div, neg_ofNat_succ]; rfl @[simp] protected theorem div_neg : ∀ a b : Int, a.div (-b) = -(a.div b) | ofNat m, 0 => show ofNat (m / 0) = -↑(m / 0) by rw [Nat.div_zero]; rfl diff --git a/src/Init/Data/Int/Lemmas.lean b/src/Init/Data/Int/Lemmas.lean index b0c98c77ca7c..ec05e8f5edcb 100644 --- a/src/Init/Data/Int/Lemmas.lean +++ b/src/Init/Data/Int/Lemmas.lean @@ -153,7 +153,7 @@ theorem subNatNat_sub (h : n ≤ m) (k : Nat) : subNatNat (m - n) k = subNatNat theorem subNatNat_add (m n k : Nat) : subNatNat (m + n) k = m + subNatNat n k := by cases n.lt_or_ge k with | inl h' => - simp [subNatNat_of_lt h', succ_pred_eq_of_pos (Nat.sub_pos_of_lt h')] + simp [subNatNat_of_lt h', sub_one_add_one_eq_of_pos (Nat.sub_pos_of_lt h')] conv => lhs; rw [← Nat.sub_add_cancel (Nat.le_of_lt h')] apply subNatNat_add_add | inr h' => simp [subNatNat_of_le h', @@ -169,12 +169,11 @@ theorem subNatNat_add_negSucc (m n k : Nat) : rw [subNatNat_sub h', Nat.add_comm] | inl h' => have h₂ : m < n + succ k := Nat.lt_of_lt_of_le h' (le_add_right _ _) - have h₃ : m ≤ n + k := le_of_succ_le_succ h₂ rw [subNatNat_of_lt h', subNatNat_of_lt h₂] - simp [Nat.add_comm] - rw [← add_succ, succ_pred_eq_of_pos (Nat.sub_pos_of_lt h'), add_succ, succ_sub h₃, - Nat.pred_succ] - rw [Nat.add_comm n, Nat.add_sub_assoc (Nat.le_of_lt h')] + simp only [pred_eq_sub_one, negSucc_add_negSucc, succ_eq_add_one, negSucc.injEq] + rw [Nat.add_right_comm, sub_one_add_one_eq_of_pos (Nat.sub_pos_of_lt h'), Nat.sub_sub, + ← Nat.add_assoc, succ_sub_succ_eq_sub, Nat.add_comm n,Nat.add_sub_assoc (Nat.le_of_lt h'), + Nat.add_comm] protected theorem add_assoc : ∀ a b c : Int, a + b + c = a + (b + c) | (m:Nat), (n:Nat), c => aux1 .. @@ -188,15 +187,15 @@ protected theorem add_assoc : ∀ a b c : Int, a + b + c = a + (b + c) | (m:Nat), -[n+1], -[k+1] => by rw [Int.add_comm, Int.add_comm m, Int.add_comm m, ← aux2, Int.add_comm -[k+1]] | -[m+1], -[n+1], -[k+1] => by - simp [add_succ, Nat.add_comm, Nat.add_left_comm, neg_ofNat_succ] + simp [Nat.add_comm, Nat.add_left_comm, Nat.add_assoc] where aux1 (m n : Nat) : ∀ c : Int, m + n + c = m + (n + c) | (k:Nat) => by simp [Nat.add_assoc] | -[k+1] => by simp [subNatNat_add] aux2 (m n k : Nat) : -[m+1] + -[n+1] + k = -[m+1] + (-[n+1] + k) := by - simp [add_succ] + simp rw [Int.add_comm, subNatNat_add_negSucc] - simp [add_succ, succ_add, Nat.add_comm] + simp [Nat.add_comm, Nat.add_left_comm, Nat.add_assoc] protected theorem add_left_comm (a b c : Int) : a + (b + c) = b + (a + c) := by rw [← Int.add_assoc, Int.add_comm a, Int.add_assoc] @@ -391,7 +390,7 @@ theorem ofNat_mul_subNatNat (m n k : Nat) : | inl h => have h' : succ m * n < succ m * k := Nat.mul_lt_mul_of_pos_left h (Nat.succ_pos m) simp [subNatNat_of_lt h, subNatNat_of_lt h'] - rw [succ_pred_eq_of_pos (Nat.sub_pos_of_lt h), ← neg_ofNat_succ, Nat.mul_sub_left_distrib, + rw [sub_one_add_one_eq_of_pos (Nat.sub_pos_of_lt h), ← neg_ofNat_succ, Nat.mul_sub_left_distrib, ← succ_pred_eq_of_pos (Nat.sub_pos_of_lt h')]; rfl | inr h => have h' : succ m * k ≤ succ m * n := Nat.mul_le_mul_left _ h @@ -406,7 +405,7 @@ theorem negSucc_mul_subNatNat (m n k : Nat) : | inl h => have h' : succ m * n < succ m * k := Nat.mul_lt_mul_of_pos_left h (Nat.succ_pos m) rw [subNatNat_of_lt h, subNatNat_of_le (Nat.le_of_lt h')] - simp [succ_pred_eq_of_pos (Nat.sub_pos_of_lt h), Nat.mul_sub_left_distrib] + simp [sub_one_add_one_eq_of_pos (Nat.sub_pos_of_lt h), Nat.mul_sub_left_distrib] | inr h => cases Nat.lt_or_ge k n with | inl h' => have h₁ : succ m * n > succ m * k := Nat.mul_lt_mul_of_pos_left h' (Nat.succ_pos m) @@ -422,12 +421,12 @@ protected theorem mul_add : ∀ a b c : Int, a * (b + c) = a * b + a * c simp [negOfNat_eq_subNatNat_zero]; rw [← subNatNat_add]; rfl | (m:Nat), -[n+1], (k:Nat) => by simp [negOfNat_eq_subNatNat_zero]; rw [Int.add_comm, ← subNatNat_add]; rfl - | (m:Nat), -[n+1], -[k+1] => by simp; rw [← Nat.left_distrib, succ_add]; rfl + | (m:Nat), -[n+1], -[k+1] => by simp [← Nat.left_distrib, Nat.add_left_comm, Nat.add_assoc] | -[m+1], (n:Nat), (k:Nat) => by simp [Nat.mul_comm]; rw [← Nat.right_distrib, Nat.mul_comm] | -[m+1], (n:Nat), -[k+1] => by simp [negOfNat_eq_subNatNat_zero]; rw [Int.add_comm, ← subNatNat_add]; rfl | -[m+1], -[n+1], (k:Nat) => by simp [negOfNat_eq_subNatNat_zero]; rw [← subNatNat_add]; rfl - | -[m+1], -[n+1], -[k+1] => by simp; rw [← Nat.left_distrib, succ_add]; rfl + | -[m+1], -[n+1], -[k+1] => by simp [← Nat.left_distrib, Nat.add_left_comm, Nat.add_assoc] protected theorem add_mul (a b c : Int) : (a + b) * c = a * c + b * c := by simp [Int.mul_comm, Int.mul_add] diff --git a/src/Init/Data/Nat/Basic.lean b/src/Init/Data/Nat/Basic.lean index 8bee45a010ef..538144874382 100644 --- a/src/Init/Data/Nat/Basic.lean +++ b/src/Init/Data/Nat/Basic.lean @@ -148,9 +148,12 @@ theorem add_succ (n m : Nat) : n + succ m = succ (n + m) := theorem add_one (n : Nat) : n + 1 = succ n := rfl -theorem succ_eq_add_one (n : Nat) : succ n = n + 1 := +@[simp] theorem succ_eq_add_one (n : Nat) : succ n = n + 1 := rfl +@[simp] theorem add_one_ne_zero (n : Nat) : n + 1 ≠ 0 := nofun +@[simp] theorem zero_ne_add_one (n : Nat) : 0 ≠ n + 1 := nofun + protected theorem add_comm : ∀ (n m : Nat), n + m = m + n | n, 0 => Eq.symm (Nat.zero_add n) | n, m+1 => by @@ -283,7 +286,7 @@ theorem succ_sub_succ (n m : Nat) : succ n - succ m = n - m := theorem sub_add_eq (a b c : Nat) : a - (b + c) = a - b - c := by induction c with | zero => simp - | succ c ih => simp [Nat.add_succ, Nat.sub_succ, ih] + | succ c ih => simp only [Nat.add_succ, Nat.sub_succ, ih] protected theorem lt_of_lt_of_le {n m k : Nat} : n < m → m ≤ k → n < k := Nat.le_trans @@ -632,8 +635,6 @@ protected theorem zero_ne_one : 0 ≠ (1 : Nat) := @[simp] theorem succ_ne_zero (n : Nat) : succ n ≠ 0 := fun h => Nat.noConfusion h -theorem add_one_ne_zero (n) : n + 1 ≠ 0 := succ_ne_zero _ - /-! # mul + order -/ theorem mul_le_mul_left {n m : Nat} (k : Nat) (h : n ≤ m) : k * n ≤ k * m := @@ -742,6 +743,11 @@ theorem succ_pred {a : Nat} (h : a ≠ 0) : a.pred.succ = a := by theorem succ_pred_eq_of_pos : ∀ {n}, 0 < n → succ (pred n) = n | _+1, _ => rfl +theorem sub_one_add_one_eq_of_pos : ∀ {n}, 0 < n → (n - 1) + 1 = n + | _+1, _ => rfl + +@[simp] theorem pred_eq_sub_one : pred n = n - 1 := rfl + /-! # sub theorems -/ theorem add_sub_self_left (a b : Nat) : (a + b) - a = b := by @@ -778,7 +784,7 @@ theorem sub_succ_lt_self (a i : Nat) (h : i < a) : a - (i + 1) < a - i := by theorem sub_ne_zero_of_lt : {a b : Nat} → a < b → b - a ≠ 0 | 0, 0, h => absurd h (Nat.lt_irrefl 0) - | 0, succ b, _ => by simp + | 0, succ b, _ => by simp only [Nat.sub_zero, ne_eq, not_false_eq_true] | succ a, 0, h => absurd h (Nat.not_lt_zero a.succ) | succ a, succ b, h => by rw [Nat.succ_sub_succ]; exact sub_ne_zero_of_lt (Nat.lt_of_succ_lt_succ h) @@ -796,7 +802,7 @@ theorem add_sub_of_le {a b : Nat} (h : a ≤ b) : a + (b - a) = b := by protected theorem add_sub_add_right (n k m : Nat) : (n + k) - (m + k) = n - m := by induction k with | zero => simp - | succ k ih => simp [add_succ, add_succ, succ_sub_succ, ih] + | succ k ih => simp [← Nat.add_assoc, ih] protected theorem add_sub_add_left (k n m : Nat) : (k + n) - (k + m) = n - m := by rw [Nat.add_comm k n, Nat.add_comm k m, Nat.add_sub_add_right] diff --git a/src/Init/Data/Nat/Bitwise/Basic.lean b/src/Init/Data/Nat/Bitwise/Basic.lean index ebf3ea1df409..f5e5a940497f 100644 --- a/src/Init/Data/Nat/Bitwise/Basic.lean +++ b/src/Init/Data/Nat/Bitwise/Basic.lean @@ -63,7 +63,7 @@ theorem shiftRight_succ (m n) : m >>> (n + 1) = (m >>> n) / 2 := rfl theorem shiftRight_add (m n : Nat) : ∀ k, m >>> (n + k) = (m >>> n) >>> k | 0 => rfl - | k + 1 => by simp [add_succ, shiftRight_add, shiftRight_succ] + | k + 1 => by simp [← Nat.add_assoc, shiftRight_add _ _ k, shiftRight_succ] theorem shiftRight_eq_div_pow (m : Nat) : ∀ n, m >>> n = m / 2 ^ n | 0 => (Nat.div_one _).symm diff --git a/src/Init/Data/Nat/Div.lean b/src/Init/Data/Nat/Div.lean index eb6be7c375cd..a8fed3caf2ae 100644 --- a/src/Init/Data/Nat/Div.lean +++ b/src/Init/Data/Nat/Div.lean @@ -293,7 +293,7 @@ theorem sub_mul_div (x n p : Nat) (h₁ : n*p ≤ x) : (x - n*p) / n = x / n - p rw [mul_succ] at h₁ exact h₁ rw [sub_succ, ← IH h₂, div_eq_sub_div h₀ h₃] - simp [add_one, Nat.pred_succ, mul_succ, Nat.sub_sub] + simp [Nat.pred_succ, mul_succ, Nat.sub_sub] theorem mul_sub_div (x n p : Nat) (h₁ : x < n*p) : (n * p - succ x) / n = p - succ (x / n) := by have npos : 0 < n := (eq_zero_or_pos _).resolve_left fun n0 => by diff --git a/src/Init/Data/Nat/Lemmas.lean b/src/Init/Data/Nat/Lemmas.lean index cdc3e696f9f6..ce83122b6572 100644 --- a/src/Init/Data/Nat/Lemmas.lean +++ b/src/Init/Data/Nat/Lemmas.lean @@ -19,7 +19,6 @@ and later these lemmas should be organised into other files more systematically. -/ namespace Nat - /-! ## add -/ protected theorem add_add_add_comm (a b c d : Nat) : (a + b) + (c + d) = (a + c) + (b + d) := by @@ -781,7 +780,7 @@ theorem shiftRight_succ_inside : ∀m n, m >>> (n+1) = (m/2) >>> n theorem shiftLeft_shiftLeft (m n : Nat) : ∀ k, (m <<< n) <<< k = m <<< (n + k) | 0 => rfl - | k + 1 => by simp [add_succ, shiftLeft_shiftLeft _ _ k, shiftLeft_succ] + | k + 1 => by simp [← Nat.add_assoc, shiftLeft_shiftLeft _ _ k, shiftLeft_succ] theorem mul_add_div {m : Nat} (m_pos : m > 0) (x y : Nat) : (m * x + y) / m = x + y / m := by match x with diff --git a/src/Init/Prelude.lean b/src/Init/Prelude.lean index fb5241f4d884..cd7f37c01f1c 100644 --- a/src/Init/Prelude.lean +++ b/src/Init/Prelude.lean @@ -1485,6 +1485,7 @@ instance [ShiftRight α] : HShiftRight α α α where hShiftRight a b := ShiftRight.shiftRight a b open HAdd (hAdd) +open HSub (hSub) open HMul (hMul) open HPow (hPow) open HAppend (hAppend) @@ -2053,7 +2054,7 @@ instance : OfNat (Fin (n+1)) i where ofNat := Fin.ofNat i ``` -/ -abbrev USize.size : Nat := Nat.succ (Nat.sub (hPow 2 System.Platform.numBits) 1) +abbrev USize.size : Nat := hAdd (hSub (hPow 2 System.Platform.numBits) 1) 1 theorem usize_size_eq : Or (Eq USize.size 4294967296) (Eq USize.size 18446744073709551616) := show Or (Eq (Nat.succ (Nat.sub (hPow 2 System.Platform.numBits) 1)) 4294967296) (Eq (Nat.succ (Nat.sub (hPow 2 System.Platform.numBits) 1)) 18446744073709551616) from diff --git a/tests/lean/guessLexFailures.lean.expected.out b/tests/lean/guessLexFailures.lean.expected.out index ca6d94d2d1eb..7ca2512b31ff 100644 --- a/tests/lean/guessLexFailures.lean.expected.out +++ b/tests/lean/guessLexFailures.lean.expected.out @@ -146,4 +146,4 @@ guessLexFailures.lean:131:14-131:31: error: failed to prove termination, possibl - Use `termination_by` to specify a different well-founded relation - Use `decreasing_by` to specify your own tactic for discharging this kind of goal m n✝ n : Nat -⊢ n < Nat.succ n✝ +⊢ n < n✝ + 1 diff --git a/tests/lean/matchMultAlt.lean b/tests/lean/matchMultAlt.lean index c61cfe6f3470..ac78732122f5 100644 --- a/tests/lean/matchMultAlt.lean +++ b/tests/lean/matchMultAlt.lean @@ -24,7 +24,7 @@ def h (x : Nat) : Nat := example : h x > 0 := by match x with | 0 | 1 => decide - | x + 2 => simp [h, Nat.add_succ] + | x + 2 => simp [h] inductive StrOrNum where | S (s : String) diff --git a/tests/lean/run/ac_rfl.lean b/tests/lean/run/ac_rfl.lean index eb1bc983ac65..adebc67845a7 100644 --- a/tests/lean/run/ac_rfl.lean +++ b/tests/lean/run/ac_rfl.lean @@ -12,7 +12,7 @@ instance : LawfulCommIdentity HMul.hMul 1 where right_id := Nat.mul_one ⟨Nat.le_of_succ_le_succ, Nat.succ_le_succ⟩ @[simp] theorem add_le_add_right_iff {x y z : Nat} : x + z ≤ y + z ↔ x ≤ y := by - induction z <;> simp_all [Nat.add_succ] + induction z <;> simp_all [← Nat.add_assoc] set_option linter.unusedVariables false in theorem le_ext : ∀ {x y : Nat} (h : ∀ z, z ≤ x ↔ z ≤ y), x = y diff --git a/tests/lean/run/eqnsAtSimp.lean b/tests/lean/run/eqnsAtSimp.lean index 6f22893a0b8d..e6870e8d9194 100644 --- a/tests/lean/run/eqnsAtSimp.lean +++ b/tests/lean/run/eqnsAtSimp.lean @@ -15,4 +15,4 @@ end theorem isEven_double (x : Nat) : isEven (2 * x) = true := by induction x with | zero => simp [isEven] - | succ x ih => simp [Nat.mul_succ, Nat.add_succ, isEven, isOdd, ih] + | succ x ih => simp [Nat.mul_succ, isEven, isOdd, ih] diff --git a/tests/lean/run/eqnsAtSimp2.lean b/tests/lean/run/eqnsAtSimp2.lean index 9f2f33d2791c..6676b796fe23 100644 --- a/tests/lean/run/eqnsAtSimp2.lean +++ b/tests/lean/run/eqnsAtSimp2.lean @@ -13,7 +13,7 @@ end theorem isEven_double (x : Nat) : isEven (2 * x) = true := by induction x with | zero => simp - | succ x ih => simp [Nat.mul_succ, Nat.add_succ, ih] + | succ x ih => simp [Nat.mul_succ, ih] def f (x : Nat) : Nat := match x with diff --git a/tests/lean/run/simp3.lean b/tests/lean/run/simp3.lean index 3e59f0bfec6c..382a4e2cdc97 100644 --- a/tests/lean/run/simp3.lean +++ b/tests/lean/run/simp3.lean @@ -7,7 +7,6 @@ theorem ex2 (x : Nat) : f ((fun a => f (f a)) x) = x := by simp opaque g (x : Nat) : Nat @[simp] axiom gEq (x : Nat) : g x = match x with | 0 => 1 | x+1 => 2 -@[simp] theorem add1 (x : Nat) : x + 1 = x.succ := rfl theorem ex3 (x : Nat) : g (x + 1) = 2 := by simp diff --git a/tests/lean/run/simpRwBug.lean b/tests/lean/run/simpRwBug.lean index bf2d440d40c1..91e62e5d8f9d 100644 --- a/tests/lean/run/simpRwBug.lean +++ b/tests/lean/run/simpRwBug.lean @@ -1,4 +1,4 @@ open Nat theorem add_assoc (m n k : Nat) : m + n + k = m + (n + k) := - Nat.recOn (motive := fun k => m + n + k = m + (n + k)) k rfl (fun k ih => by simp [Nat.add_succ, ih]; done) + Nat.recOn (motive := fun k => m + n + k = m + (n + k)) k rfl (fun k ih => by simp [Nat.add_assoc, ih]; done) diff --git a/tests/lean/run/where1.lean b/tests/lean/run/where1.lean index 116141556004..45cdd4e2e52b 100644 --- a/tests/lean/run/where1.lean +++ b/tests/lean/run/where1.lean @@ -23,7 +23,7 @@ where | nil => simp [reverse.loop] | cons a as ih => show (reverse.loop as (a::bs)).length = (a :: as).length + bs.length - simp [ih, Nat.add_succ, Nat.succ_add] + simp [ih, Nat.add_assoc, Nat.succ_add] def h : Nat -> Nat | 0 => g 0 From 20038140853cf121d401d281bd30bbcee871797d Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Wed, 13 Mar 2024 00:56:27 -0700 Subject: [PATCH 20/32] chore: rename automatically generated equational theorems (#3661) cc @nomeata --- RELEASES.md | 27 +++++++++++ src/Lean/Elab/PreDefinition/Eqns.lean | 2 +- .../Elab/PreDefinition/Structural/Eqns.lean | 2 +- src/Lean/Elab/PreDefinition/WF/Eqns.lean | 2 +- src/Lean/Meta/Tactic/Simp/Types.lean | 2 +- tests/lean/1026.lean.expected.out | 9 ---- tests/lean/1074a.lean.expected.out | 4 -- tests/lean/974.lean | 20 -------- tests/lean/974.lean.expected.out | 9 ---- tests/lean/986.lean | 4 -- tests/lean/986.lean.expected.out | 9 ---- tests/lean/eqValue.lean | 11 ----- tests/lean/eqValue.lean.expected.out | 4 -- tests/lean/heapSort.lean.expected.out | 17 ------- tests/lean/namePatEqThm.lean | 15 ------ tests/lean/namePatEqThm.lean.expected.out | 8 ---- tests/lean/{ => run}/1026.lean | 12 ++++- tests/lean/{ => run}/1074a.lean | 9 +++- tests/lean/run/1080.lean | 4 +- tests/lean/run/974.lean | 40 ++++++++++++++++ tests/lean/run/986.lean | 19 ++++++++ tests/lean/run/byteSliceIssue.lean | 2 +- tests/lean/run/eqValue.lean | 21 +++++++++ tests/lean/run/eqnsAtSimp3.lean | 4 +- tests/lean/{ => run}/heapSort.lean | 13 +++++- tests/lean/run/instEtaIssue.lean | 4 +- tests/lean/run/match_lit_fin_cover.lean | 20 ++++---- tests/lean/run/match_lit_issues.lean | 20 ++++---- tests/lean/run/namePatEqThm.lean | 35 ++++++++++++++ tests/lean/run/nestedIssueMatch.lean | 4 +- tests/lean/run/nestedWF.lean | 14 +++--- tests/lean/run/overAndPartialAppsAtWF.lean | 2 +- tests/lean/run/printEqns.lean | 14 +++--- tests/lean/run/robinson.lean | 8 ++-- tests/lean/run/rossel1.lean | 6 +-- tests/lean/run/simpAtDefIssue.lean | 8 ++-- tests/lean/run/simp_eqn_bug.lean | 6 +-- tests/lean/run/splitIssue.lean | 8 ++-- tests/lean/run/splitList.lean | 16 +++---- tests/lean/run/streamEqIssue.lean | 6 +-- .../structEqns.lean} | 23 +++++++++- tests/lean/run/structuralEqns.lean | 46 +++++++++---------- tests/lean/run/structuralEqns2.lean | 12 ++--- tests/lean/run/structuralEqns3.lean | 6 +-- tests/lean/run/wfEqns1.lean | 6 +-- tests/lean/run/wfEqns2.lean | 12 ++--- tests/lean/run/wfEqns3.lean | 4 +- tests/lean/run/wfEqns4.lean | 12 ++--- tests/lean/structuralEqns.lean.expected.out | 15 ------ 49 files changed, 321 insertions(+), 255 deletions(-) delete mode 100644 tests/lean/1026.lean.expected.out delete mode 100644 tests/lean/1074a.lean.expected.out delete mode 100644 tests/lean/974.lean delete mode 100644 tests/lean/974.lean.expected.out delete mode 100644 tests/lean/986.lean delete mode 100644 tests/lean/986.lean.expected.out delete mode 100644 tests/lean/eqValue.lean delete mode 100644 tests/lean/eqValue.lean.expected.out delete mode 100644 tests/lean/heapSort.lean.expected.out delete mode 100644 tests/lean/namePatEqThm.lean delete mode 100644 tests/lean/namePatEqThm.lean.expected.out rename tests/lean/{ => run}/1026.lean (55%) rename tests/lean/{ => run}/1074a.lean (69%) create mode 100644 tests/lean/run/974.lean create mode 100644 tests/lean/run/eqValue.lean rename tests/lean/{ => run}/heapSort.lean (95%) create mode 100644 tests/lean/run/namePatEqThm.lean rename tests/lean/{structuralEqns.lean => run/structEqns.lean} (55%) delete mode 100644 tests/lean/structuralEqns.lean.expected.out diff --git a/RELEASES.md b/RELEASES.md index 18b675380136..e6c7e073ddc0 100644 --- a/RELEASES.md +++ b/RELEASES.md @@ -31,6 +31,33 @@ v4.8.0 (development in progress) (x x : Nat) : motive x x ``` +Breaking changes: + +* Automatically generated equational theorems are now named using suffix `.eq_` instead of `._eq_`, and `.def` instead of `._unfold`. Example: +``` +def fact : Nat → Nat + | 0 => 1 + | n+1 => (n+1) * fact n + +theorem ex : fact 0 = 1 := by unfold fact; decide + +#check fact.eq_1 +-- fact.eq_1 : fact 0 = 1 + +#check fact.eq_2 +-- fact.eq_2 (n : Nat) : fact (Nat.succ n) = (n + 1) * fact n + +#check fact.def +/- +fact.def : + ∀ (x : Nat), + fact x = + match x with + | 0 => 1 + | Nat.succ n => (n + 1) * fact n +-/ +``` + v4.7.0 --------- diff --git a/src/Lean/Elab/PreDefinition/Eqns.lean b/src/Lean/Elab/PreDefinition/Eqns.lean index 2a2eadd1b479..e0cc1c4f05cc 100644 --- a/src/Lean/Elab/PreDefinition/Eqns.lean +++ b/src/Lean/Elab/PreDefinition/Eqns.lean @@ -380,7 +380,7 @@ def mkUnfoldEq (declName : Name) (info : EqnInfoCore) : MetaM Name := withLCtx { mkUnfoldProof declName goal.mvarId! let type ← mkForallFVars xs type let value ← mkLambdaFVars xs (← instantiateMVars goal) - let name := baseName ++ `_unfold + let name := baseName ++ `def addDecl <| Declaration.thmDecl { name, type, value levelParams := info.levelParams diff --git a/src/Lean/Elab/PreDefinition/Structural/Eqns.lean b/src/Lean/Elab/PreDefinition/Structural/Eqns.lean index 6738dddb6b25..5ee888c748da 100644 --- a/src/Lean/Elab/PreDefinition/Structural/Eqns.lean +++ b/src/Lean/Elab/PreDefinition/Structural/Eqns.lean @@ -68,7 +68,7 @@ def mkEqns (info : EqnInfo) : MetaM (Array Name) := for i in [: eqnTypes.size] do let type := eqnTypes[i]! trace[Elab.definition.structural.eqns] "{eqnTypes[i]!}" - let name := baseName ++ (`_eq).appendIndexAfter (i+1) + let name := baseName ++ (`eq).appendIndexAfter (i+1) thmNames := thmNames.push name let value ← mkProof info.declName type let (type, value) ← removeUnusedEqnHypotheses type value diff --git a/src/Lean/Elab/PreDefinition/WF/Eqns.lean b/src/Lean/Elab/PreDefinition/WF/Eqns.lean index e0c71d28fae2..f4e6b35808ea 100644 --- a/src/Lean/Elab/PreDefinition/WF/Eqns.lean +++ b/src/Lean/Elab/PreDefinition/WF/Eqns.lean @@ -117,7 +117,7 @@ def mkEqns (declName : Name) (info : EqnInfo) : MetaM (Array Name) := for i in [: eqnTypes.size] do let type := eqnTypes[i]! trace[Elab.definition.wf.eqns] "{eqnTypes[i]!}" - let name := baseName ++ (`_eq).appendIndexAfter (i+1) + let name := baseName ++ (`eq).appendIndexAfter (i+1) thmNames := thmNames.push name let value ← mkProof declName type let (type, value) ← removeUnusedEqnHypotheses type value diff --git a/src/Lean/Meta/Tactic/Simp/Types.lean b/src/Lean/Meta/Tactic/Simp/Types.lean index 659effcec3ef..2540eb529f6b 100644 --- a/src/Lean/Meta/Tactic/Simp/Types.lean +++ b/src/Lean/Meta/Tactic/Simp/Types.lean @@ -285,7 +285,7 @@ def getSimpCongrTheorems : SimpM SimpCongrTheorems := def recordSimpTheorem (thmId : Origin) : SimpM Unit := do /- - If `thmId` is an equational theorem (e.g., `foo._eq_1`), we should record `foo` instead. + If `thmId` is an equational theorem (e.g., `foo.eq_1`), we should record `foo` instead. See issue #3547. -/ let thmId ← match thmId with diff --git a/tests/lean/1026.lean.expected.out b/tests/lean/1026.lean.expected.out deleted file mode 100644 index 64966ecc9ad9..000000000000 --- a/tests/lean/1026.lean.expected.out +++ /dev/null @@ -1,9 +0,0 @@ -1026.lean:1:4-1:7: warning: declaration uses 'sorry' -1026.lean:9:8-9:10: warning: declaration uses 'sorry' -foo._unfold (n : Nat) : - foo n = - if n = 0 then 0 - else - let x := n - 1; - let_fun this := foo.proof_4; - foo x diff --git a/tests/lean/1074a.lean.expected.out b/tests/lean/1074a.lean.expected.out deleted file mode 100644 index f5b7e7683e53..000000000000 --- a/tests/lean/1074a.lean.expected.out +++ /dev/null @@ -1,4 +0,0 @@ -Brx.interp._eq_1 (n z : Term) (H_2 : Brx (Term.id2 n z)) : - Brx.interp H_2 = - match ⋯ with - | ⋯ => Brx.interp Hz diff --git a/tests/lean/974.lean b/tests/lean/974.lean deleted file mode 100644 index 0be5d5203b66..000000000000 --- a/tests/lean/974.lean +++ /dev/null @@ -1,20 +0,0 @@ -inductive Formula : Nat → Type u -| bot : Formula n -| imp (f₁ f₂ : Formula n ) : Formula n -| all (f : Formula (n+1)) : Formula n - -def Formula.count_quantifiers : {n:Nat} → Formula n → Nat -| _, imp f₁ f₂ => f₁.count_quantifiers + f₂.count_quantifiers -| _, all f => f.count_quantifiers + 1 -| _, _ => 0 - -attribute [simp] Formula.count_quantifiers - -#check Formula.count_quantifiers._eq_1 -#check Formula.count_quantifiers._eq_2 -#check Formula.count_quantifiers._eq_3 - -@[simp] def Formula.count_quantifiers2 : Formula n → Nat -| imp f₁ f₂ => f₁.count_quantifiers2 + f₂.count_quantifiers2 -| all f => f.count_quantifiers2 + 1 -| _ => 0 diff --git a/tests/lean/974.lean.expected.out b/tests/lean/974.lean.expected.out deleted file mode 100644 index afc0622e0ac4..000000000000 --- a/tests/lean/974.lean.expected.out +++ /dev/null @@ -1,9 +0,0 @@ -Formula.count_quantifiers._eq_1.{u_1} : - ∀ (x : Nat) (f₁ f₂ : Formula x), - Formula.count_quantifiers (Formula.imp f₁ f₂) = Formula.count_quantifiers f₁ + Formula.count_quantifiers f₂ -Formula.count_quantifiers._eq_2.{u_1} : - ∀ (x : Nat) (f : Formula (x + 1)), Formula.count_quantifiers (Formula.all f) = Formula.count_quantifiers f + 1 -Formula.count_quantifiers._eq_3.{u_1} : - ∀ (x : Nat) (x_1 : Formula x), - (∀ (f₁ f₂ : Formula x), x_1 = Formula.imp f₁ f₂ → False) → - (∀ (f : Formula (x + 1)), x_1 = Formula.all f → False) → Formula.count_quantifiers x_1 = 0 diff --git a/tests/lean/986.lean b/tests/lean/986.lean deleted file mode 100644 index b5ba6793e285..000000000000 --- a/tests/lean/986.lean +++ /dev/null @@ -1,4 +0,0 @@ -attribute [simp] Array.insertionSort.swapLoop - -#check Array.insertionSort.swapLoop._eq_1 -#check Array.insertionSort.swapLoop._eq_2 diff --git a/tests/lean/986.lean.expected.out b/tests/lean/986.lean.expected.out deleted file mode 100644 index 885bd531d552..000000000000 --- a/tests/lean/986.lean.expected.out +++ /dev/null @@ -1,9 +0,0 @@ -Array.insertionSort.swapLoop._eq_1.{u_1} {α : Type u_1} (lt : α → α → Bool) (a : Array α) (h : 0 < Array.size a) : - Array.insertionSort.swapLoop lt a 0 h = a -Array.insertionSort.swapLoop._eq_2.{u_1} {α : Type u_1} (lt : α → α → Bool) (a : Array α) (j' : Nat) - (h : Nat.succ j' < Array.size a) : - Array.insertionSort.swapLoop lt a (Nat.succ j') h = - let_fun h' := ⋯; - if lt a[Nat.succ j'] a[j'] = true then - Array.insertionSort.swapLoop lt (Array.swap a { val := Nat.succ j', isLt := h } { val := j', isLt := h' }) j' ⋯ - else a diff --git a/tests/lean/eqValue.lean b/tests/lean/eqValue.lean deleted file mode 100644 index 2ccbcc4bfc9f..000000000000 --- a/tests/lean/eqValue.lean +++ /dev/null @@ -1,11 +0,0 @@ -@[simp] def f (x : Nat) : Nat := - match x with - | 0 => 1 - | 100 => 2 - | 1000 => 3 - | x+1 => f x - -#check f._eq_1 -#check f._eq_2 -#check f._eq_3 -#check f._eq_4 diff --git a/tests/lean/eqValue.lean.expected.out b/tests/lean/eqValue.lean.expected.out deleted file mode 100644 index 05601457a815..000000000000 --- a/tests/lean/eqValue.lean.expected.out +++ /dev/null @@ -1,4 +0,0 @@ -f._eq_1 : f 0 = 1 -f._eq_2 : f 100 = 2 -f._eq_3 : f 1000 = 3 -f._eq_4 (x_2 : Nat) (x_3 : x_2 = 99 → False) (x_4 : x_2 = 999 → False) : f (Nat.succ x_2) = f x_2 diff --git a/tests/lean/heapSort.lean.expected.out b/tests/lean/heapSort.lean.expected.out deleted file mode 100644 index 9e0ab4f588f0..000000000000 --- a/tests/lean/heapSort.lean.expected.out +++ /dev/null @@ -1,17 +0,0 @@ -heapSort.lean:15:4-15:15: warning: declaration uses 'sorry' -heapSort.lean:15:4-15:15: warning: declaration uses 'sorry' -heapSort.lean:15:4-15:15: warning: declaration uses 'sorry' -heapSort.lean:43:4-43:10: warning: declaration uses 'sorry' -heapSort.lean:58:4-58:13: warning: declaration uses 'sorry' -heapSort.lean:58:4-58:13: warning: declaration uses 'sorry' -heapSort.lean:58:4-58:13: warning: declaration uses 'sorry' -heapSort.lean:102:4-102:13: warning: declaration uses 'sorry' -heapSort.lean:102:4-102:13: warning: declaration uses 'sorry' -Array.heapSort.loop._eq_1.{u_1} {α : Type u_1} (lt : α → α → Bool) (a : BinaryHeap α fun y x => lt x y) - (out : Array α) : - Array.heapSort.loop lt a out = - match e : BinaryHeap.max a with - | none => out - | some x => - let_fun this := ⋯; - Array.heapSort.loop lt (BinaryHeap.popMax a) (Array.push out x) diff --git a/tests/lean/namePatEqThm.lean b/tests/lean/namePatEqThm.lean deleted file mode 100644 index 18c118508212..000000000000 --- a/tests/lean/namePatEqThm.lean +++ /dev/null @@ -1,15 +0,0 @@ -@[simp] def iota : Nat → List Nat - | 0 => [] - | m@(n+1) => m :: iota n - -#check iota._eq_1 -#check iota._eq_2 - -@[simp] def f : List Nat → List Nat × List Nat - | xs@(x :: ys@(y :: [])) => (xs, ys) - | xs@(x :: ys@(y :: zs)) => f zs - | _ => ([], []) - -#check f._eq_1 -#check f._eq_2 -#check f._eq_3 diff --git a/tests/lean/namePatEqThm.lean.expected.out b/tests/lean/namePatEqThm.lean.expected.out deleted file mode 100644 index 73e55b147dc2..000000000000 --- a/tests/lean/namePatEqThm.lean.expected.out +++ /dev/null @@ -1,8 +0,0 @@ -iota._eq_1 : iota 0 = [] -iota._eq_2 (n : Nat) : iota (Nat.succ n) = Nat.succ n :: iota n -f._eq_1 (x_1 y : Nat) : f [x_1, y] = ([x_1, y], [y]) -f._eq_2 (x_1 y : Nat) (zs : List Nat) (x_2 : zs = [] → False) : f (x_1 :: y :: zs) = f zs -f._eq_3 : - ∀ (x : List Nat), - (∀ (x_1 y : Nat), x = [x_1, y] → False) → - (∀ (x_1 y : Nat) (zs : List Nat), x = x_1 :: y :: zs → False) → f x = ([], []) diff --git a/tests/lean/1026.lean b/tests/lean/run/1026.lean similarity index 55% rename from tests/lean/1026.lean rename to tests/lean/run/1026.lean index afe792b15679..ec06a943efb4 100644 --- a/tests/lean/1026.lean +++ b/tests/lean/run/1026.lean @@ -10,4 +10,14 @@ theorem ex : foo 0 = 0 := by unfold foo sorry -#check foo._unfold +/-- +info: foo.def (n : Nat) : + foo n = + if n = 0 then 0 + else + let x := n - 1; + let_fun this := foo.proof_4; + foo x +-/ +#guard_msgs in +#check foo.def diff --git a/tests/lean/1074a.lean b/tests/lean/run/1074a.lean similarity index 69% rename from tests/lean/1074a.lean rename to tests/lean/run/1074a.lean index 49451e8026e8..3e5e1c9d9e85 100644 --- a/tests/lean/1074a.lean +++ b/tests/lean/run/1074a.lean @@ -18,4 +18,11 @@ def Brx.interp_nil (H: Brx a): H.interp = H.interp rfl } -#check Brx.interp._eq_1 +/-- +info: Brx.interp.eq_1 (n z : Term) (H_2 : Brx (Term.id2 n z)) : + Brx.interp H_2 = + match ⋯ with + | ⋯ => Brx.interp Hz +-/ +#guard_msgs in +#check Brx.interp.eq_1 diff --git a/tests/lean/run/1080.lean b/tests/lean/run/1080.lean index e4132a9d8621..265be5849b8a 100644 --- a/tests/lean/run/1080.lean +++ b/tests/lean/run/1080.lean @@ -17,8 +17,8 @@ def Bar.check: Bar f → Prop attribute [simp] Bar.check -#check Bar.check._eq_1 -#check Bar.check._eq_2 +#check Bar.check.eq_1 +#check Bar.check.eq_2 #check Bar.check.match_1.eq_1 #check Bar.check.match_1.eq_2 #check Bar.check.match_1.splitter diff --git a/tests/lean/run/974.lean b/tests/lean/run/974.lean new file mode 100644 index 000000000000..0077bac03a51 --- /dev/null +++ b/tests/lean/run/974.lean @@ -0,0 +1,40 @@ +inductive Formula : Nat → Type u where + | bot : Formula n + | imp (f₁ f₂ : Formula n ) : Formula n + | all (f : Formula (n+1)) : Formula n + +def Formula.count_quantifiers : {n:Nat} → Formula n → Nat + | _, imp f₁ f₂ => f₁.count_quantifiers + f₂.count_quantifiers + | _, all f => f.count_quantifiers + 1 + | _, _ => 0 + +attribute [simp] Formula.count_quantifiers + +/-- +info: Formula.count_quantifiers.eq_1.{u_1} : + ∀ (x : Nat) (f₁ f₂ : Formula x), + Formula.count_quantifiers (Formula.imp f₁ f₂) = Formula.count_quantifiers f₁ + Formula.count_quantifiers f₂ +-/ +#guard_msgs in +#check Formula.count_quantifiers.eq_1 + +/-- +info: Formula.count_quantifiers.eq_2.{u_1} : + ∀ (x : Nat) (f : Formula (x + 1)), Formula.count_quantifiers (Formula.all f) = Formula.count_quantifiers f + 1 +-/ +#guard_msgs in +#check Formula.count_quantifiers.eq_2 + +/-- +info: Formula.count_quantifiers.eq_3.{u_1} : + ∀ (x : Nat) (x_1 : Formula x), + (∀ (f₁ f₂ : Formula x), x_1 = Formula.imp f₁ f₂ → False) → + (∀ (f : Formula (x + 1)), x_1 = Formula.all f → False) → Formula.count_quantifiers x_1 = 0 +-/ +#guard_msgs in +#check Formula.count_quantifiers.eq_3 + +@[simp] def Formula.count_quantifiers2 : Formula n → Nat + | imp f₁ f₂ => f₁.count_quantifiers2 + f₂.count_quantifiers2 + | all f => f.count_quantifiers2 + 1 + | _ => 0 diff --git a/tests/lean/run/986.lean b/tests/lean/run/986.lean index 1448c57b9799..1d7dc33fe1b8 100644 --- a/tests/lean/run/986.lean +++ b/tests/lean/run/986.lean @@ -1 +1,20 @@ attribute [simp] Array.insertionSort.swapLoop + +/-- +info: Array.insertionSort.swapLoop.eq_1.{u_1} {α : Type u_1} (lt : α → α → Bool) (a : Array α) (h : 0 < Array.size a) : + Array.insertionSort.swapLoop lt a 0 h = a +-/ +#guard_msgs in +#check Array.insertionSort.swapLoop.eq_1 + +/-- +info: Array.insertionSort.swapLoop.eq_2.{u_1} {α : Type u_1} (lt : α → α → Bool) (a : Array α) (j' : Nat) + (h : Nat.succ j' < Array.size a) : + Array.insertionSort.swapLoop lt a (Nat.succ j') h = + let_fun h' := ⋯; + if lt a[Nat.succ j'] a[j'] = true then + Array.insertionSort.swapLoop lt (Array.swap a { val := Nat.succ j', isLt := h } { val := j', isLt := h' }) j' ⋯ + else a +-/ +#guard_msgs in +#check Array.insertionSort.swapLoop.eq_2 diff --git a/tests/lean/run/byteSliceIssue.lean b/tests/lean/run/byteSliceIssue.lean index 76147a5c2f28..4d2388365857 100644 --- a/tests/lean/run/byteSliceIssue.lean +++ b/tests/lean/run/byteSliceIssue.lean @@ -42,4 +42,4 @@ def forIn.loop [Monad m] (f : UInt8 → β → m (ForInStep β)) termination_by _end - i attribute [simp] ByteSlice.forIn.loop -#check @ByteSlice.forIn.loop._eq_1 +#check @ByteSlice.forIn.loop.eq_1 diff --git a/tests/lean/run/eqValue.lean b/tests/lean/run/eqValue.lean new file mode 100644 index 000000000000..1e62bc1fabac --- /dev/null +++ b/tests/lean/run/eqValue.lean @@ -0,0 +1,21 @@ +@[simp] def f (x : Nat) : Nat := + match x with + | 0 => 1 + | 100 => 2 + | 1000 => 3 + | x+1 => f x + +/-- info: f.eq_1 : f 0 = 1 -/ +#guard_msgs in +#check f.eq_1 +/-- info: f.eq_2 : f 100 = 2 -/ +#guard_msgs in +#check f.eq_2 +/-- info: f.eq_3 : f 1000 = 3 -/ +#guard_msgs in +#check f.eq_3 +/-- +info: f.eq_4 (x_2 : Nat) (x_3 : x_2 = 99 → False) (x_4 : x_2 = 999 → False) : f (Nat.succ x_2) = f x_2 +-/ +#guard_msgs in +#check f.eq_4 diff --git a/tests/lean/run/eqnsAtSimp3.lean b/tests/lean/run/eqnsAtSimp3.lean index d6341123898f..f38d5c73d5a1 100644 --- a/tests/lean/run/eqnsAtSimp3.lean +++ b/tests/lean/run/eqnsAtSimp3.lean @@ -38,7 +38,7 @@ theorem ex3 (x : Nat) (y : Nat) (h : y = 5 → False) : ∃ z, f (x+1) y = 2 * z | x+1, y, z => 2 * f2 x y z -#check f2._eq_4 +#check f2.eq_4 theorem ex4 (x y z : Nat) (h : y = 5 → z = 6 → False) : ∃ w, f2 (x+1) y z = 2 * w := by simp [f2, h] @@ -64,7 +64,7 @@ theorem ex6 (x y z : Nat) (h2 : z ≠ 6) : ∃ w, f2 (x+1) y z = 2 * w := by | x+1, 6, 4 => 3 * f3 x 0 1 | x+1, y, z => 2 * f3 x y z -#check f3._eq_5 +#check f3.eq_5 theorem ex7 (x y z : Nat) (h2 : z ≠ 6) (h3 : y ≠ 6) : ∃ w, f3 (x+1) y z = 2 * w := by simp [f3, h2, h3] diff --git a/tests/lean/heapSort.lean b/tests/lean/run/heapSort.lean similarity index 95% rename from tests/lean/heapSort.lean rename to tests/lean/run/heapSort.lean index f252f73bf604..58f7f8157175 100644 --- a/tests/lean/heapSort.lean +++ b/tests/lean/run/heapSort.lean @@ -173,6 +173,17 @@ def Array.toBinaryHeap (lt : α → α → Bool) (a : Array α) : BinaryHeap α loop (a.toBinaryHeap gt) #[] attribute [simp] Array.heapSort.loop -#check Array.heapSort.loop._eq_1 + +/-- +info: Array.heapSort.loop.eq_1.{u_1} {α : Type u_1} (lt : α → α → Bool) (a : BinaryHeap α fun y x => lt x y) (out : Array α) : + Array.heapSort.loop lt a out = + match e : BinaryHeap.max a with + | none => out + | some x => + let_fun this := ⋯; + Array.heapSort.loop lt (BinaryHeap.popMax a) (Array.push out x) +-/ +#guard_msgs in +#check Array.heapSort.loop.eq_1 attribute [simp] BinaryHeap.heapifyDown diff --git a/tests/lean/run/instEtaIssue.lean b/tests/lean/run/instEtaIssue.lean index 27f18680feb2..95414a15bcff 100644 --- a/tests/lean/run/instEtaIssue.lean +++ b/tests/lean/run/instEtaIssue.lean @@ -8,12 +8,12 @@ attribute [simp] filter set_option pp.explicit true /-- -info: filter._eq_2.{u_1} {α : Type u_1} (p : α → Prop) [@DecidablePred α p] (x : α) (xs' : List α) : +info: filter.eq_2.{u_1} {α : Type u_1} (p : α → Prop) [@DecidablePred α p] (x : α) (xs' : List α) : @Eq (List α) (@filter α p inst✝ (@List.cons α x xs')) (@ite (List α) (p x) (inst✝ x) (@List.cons α x (@filter α p inst✝ xs')) (@filter α p inst✝ xs')) -/ #guard_msgs in -#check filter._eq_2 -- We should not have terms of the form `@filter α p (fun x => inst✝ x) xs'` +#check filter.eq_2 -- We should not have terms of the form `@filter α p (fun x => inst✝ x) xs'` def filter_length (p : α → Prop) [DecidablePred p] : (filter p xs).length ≤ xs.length := by diff --git a/tests/lean/run/match_lit_fin_cover.lean b/tests/lean/run/match_lit_fin_cover.lean index ed219a99cace..53353cc4d541 100644 --- a/tests/lean/run/match_lit_fin_cover.lean +++ b/tests/lean/run/match_lit_fin_cover.lean @@ -17,16 +17,16 @@ def boo (x : Fin 3) : Nat := | 2, y+1 => bla x y + 1 /-- -info: bla._eq_1 (y : Nat) : bla 0 y = 10 +info: bla.eq_1 (y : Nat) : bla 0 y = 10 -/ #guard_msgs in -#check bla._eq_1 +#check bla.eq_1 /-- -info: bla._eq_4 (y_2 : Nat) : bla 2 (Nat.succ y_2) = bla 2 y_2 + 1 +info: bla.eq_4 (y_2 : Nat) : bla 2 (Nat.succ y_2) = bla 2 y_2 + 1 -/ #guard_msgs in -#check bla._eq_4 +#check bla.eq_4 open BitVec @@ -56,19 +56,19 @@ def foo' (x : BitVec 3) (y : Nat) : Nat := attribute [simp] foo' /-- -info: foo'._eq_1 (y : Nat) : foo' (0#3) y = 7 +info: foo'.eq_1 (y : Nat) : foo' (0#3) y = 7 -/ #guard_msgs in -#check foo'._eq_1 +#check foo'.eq_1 /-- -info: foo'._eq_2 (y : Nat) : foo' (1#3) y = 6 +info: foo'.eq_2 (y : Nat) : foo' (1#3) y = 6 -/ #guard_msgs in -#check foo'._eq_2 +#check foo'.eq_2 /-- -info: foo'._eq_9 (y_2 : Nat) : foo' (7#3) (Nat.succ y_2) = foo' 7 y_2 + 1 +info: foo'.eq_9 (y_2 : Nat) : foo' (7#3) (Nat.succ y_2) = foo' 7 y_2 + 1 -/ #guard_msgs in -#check foo'._eq_9 +#check foo'.eq_9 diff --git a/tests/lean/run/match_lit_issues.lean b/tests/lean/run/match_lit_issues.lean index 14b434e38e16..7df72feb775b 100644 --- a/tests/lean/run/match_lit_issues.lean +++ b/tests/lean/run/match_lit_issues.lean @@ -4,8 +4,8 @@ | _, 0 => b+1 | _, a+1 => f1 (i-1) a (b*2) -#check f1._eq_1 -#check f1._eq_2 +#check f1.eq_1 +#check f1.eq_2 example : f1 (-1) a b = b := by simp -- should work example : f1 (-2) 0 b = b+1 := by simp @@ -29,8 +29,8 @@ example (h : c ≠ 'a') : f2 c (a+1) b = f2 c a (b*2) := by simp | _, 0 => b+1 | _, a+1 => f3 (i+1) a (b*2) -#check f3._eq_1 -#check f3._eq_2 +#check f3.eq_1 +#check f3.eq_2 example : f3 2 a b = b := by simp -- should work example : f3 3 0 b = b+1 := by simp @@ -43,8 +43,8 @@ example (h : i ≠ 2) : f3 i (a+1) b = f3 (i+1) a (b*2) := by simp; done -- shou | _, 0 => b+1 | _, a+1 => f4 (i+1) a (b*2) -#check f4._eq_1 -#check f4._eq_2 +#check f4.eq_1 +#check f4.eq_2 example : f4 2 a b = b := by simp -- should work example : f4 3 0 b = b+1 := by simp @@ -57,8 +57,8 @@ example (h : i ≠ 2) : f4 i (a+1) b = f4 (i+1) a (b*2) := by simp -- should wor | _, 0 => b+1 | _, a+1 => f5 (i+1) a (b*2) -#check f5._eq_1 -#check f5._eq_2 +#check f5.eq_1 +#check f5.eq_2 open BitVec @@ -76,8 +76,8 @@ example (h : i ≠ 2#8) : f5 i (a+1) b = f5 (i+1) a (b*2) := by simp -- should w | _, 0 => b+1 | _, a+1 => f6 (i+1) a (b*2) -#check f6._eq_1 -#check f6._eq_2 +#check f6.eq_1 +#check f6.eq_2 example : f6 2#8 a b = b := by simp -- should work example : f6 2#8 a b = b := by simp -- should work diff --git a/tests/lean/run/namePatEqThm.lean b/tests/lean/run/namePatEqThm.lean new file mode 100644 index 000000000000..8a627ebeeea7 --- /dev/null +++ b/tests/lean/run/namePatEqThm.lean @@ -0,0 +1,35 @@ +@[simp] def iota : Nat → List Nat + | 0 => [] + | m@(n+1) => m :: iota n + +/-- info: iota.eq_1 : iota 0 = [] -/ +#guard_msgs in +#check iota.eq_1 + +/-- info: iota.eq_2 (n : Nat) : iota (Nat.succ n) = Nat.succ n :: iota n -/ +#guard_msgs in +#check iota.eq_2 + +@[simp] def f : List Nat → List Nat × List Nat + | xs@(x :: ys@(y :: [])) => (xs, ys) + | xs@(x :: ys@(y :: zs)) => f zs + | _ => ([], []) + +/-- info: f.eq_1 (x_1 y : Nat) : f [x_1, y] = ([x_1, y], [y]) -/ +#guard_msgs in +#check f.eq_1 + +/-- +info: f.eq_2 (x_1 y : Nat) (zs : List Nat) (x_2 : zs = [] → False) : f (x_1 :: y :: zs) = f zs +-/ +#guard_msgs in +#check f.eq_2 + +/-- +info: f.eq_3 : + ∀ (x : List Nat), + (∀ (x_1 y : Nat), x = [x_1, y] → False) → + (∀ (x_1 y : Nat) (zs : List Nat), x = x_1 :: y :: zs → False) → f x = ([], []) +-/ +#guard_msgs in +#check f.eq_3 diff --git a/tests/lean/run/nestedIssueMatch.lean b/tests/lean/run/nestedIssueMatch.lean index 5b0ac54b41af..a1eccc46b366 100644 --- a/tests/lean/run/nestedIssueMatch.lean +++ b/tests/lean/run/nestedIssueMatch.lean @@ -9,8 +9,8 @@ decreasing_by all_goals sorry attribute [simp] g -#check g._eq_1 -#check g._eq_2 +#check g.eq_1 +#check g.eq_2 theorem ex3 : g (n + 1) = match g n with | 0 => 0 diff --git a/tests/lean/run/nestedWF.lean b/tests/lean/run/nestedWF.lean index 8caabf2b7d26..7fea4e10d81e 100644 --- a/tests/lean/run/nestedWF.lean +++ b/tests/lean/run/nestedWF.lean @@ -28,12 +28,12 @@ attribute [simp] g attribute [simp] h attribute [simp] f -#check g._eq_1 -#check g._eq_2 +#check g.eq_1 +#check g.eq_2 -#check h._eq_1 +#check h.eq_1 -#check f._eq_1 +#check f.eq_1 end Ex1 @@ -51,14 +51,14 @@ decreasing_by all_goals sorry theorem ex1 : g 0 = 0 := by rw [g] -#check g._eq_1 -#check g._eq_2 +#check g.eq_1 +#check g.eq_2 theorem ex2 : g 0 = 0 := by unfold g simp -#check g._unfold +#check g.def end Ex2 diff --git a/tests/lean/run/overAndPartialAppsAtWF.lean b/tests/lean/run/overAndPartialAppsAtWF.lean index 6ff6a3269a2c..72110a402a72 100644 --- a/tests/lean/run/overAndPartialAppsAtWF.lean +++ b/tests/lean/run/overAndPartialAppsAtWF.lean @@ -19,4 +19,4 @@ termination_by _ _ _ => a.size - i (· + y) termination_by x -#check f._eq_1 +#check f.eq_1 diff --git a/tests/lean/run/printEqns.lean b/tests/lean/run/printEqns.lean index c0a0d06c0b8d..887c6ad555e6 100644 --- a/tests/lean/run/printEqns.lean +++ b/tests/lean/run/printEqns.lean @@ -1,7 +1,7 @@ /-- info: equations: -private theorem List.append._eq_1.{u} : ∀ {α : Type u} (x : List α), List.append [] x = x -private theorem List.append._eq_2.{u} : ∀ {α : Type u} (x : List α) (a : α) (l : List α), +private theorem List.append.eq_1.{u} : ∀ {α : Type u} (x : List α), List.append [] x = x +private theorem List.append.eq_2.{u} : ∀ {α : Type u} (x : List α) (a : α) (l : List α), List.append (a :: l) x = a :: List.append l x -/ #guard_msgs in @@ -9,8 +9,8 @@ private theorem List.append._eq_2.{u} : ∀ {α : Type u} (x : List α) (a : α) /-- info: equations: -private theorem List.append._eq_1.{u} : ∀ {α : Type u} (x : List α), List.append [] x = x -private theorem List.append._eq_2.{u} : ∀ {α : Type u} (x : List α) (a : α) (l : List α), +private theorem List.append.eq_1.{u} : ∀ {α : Type u} (x : List α), List.append [] x = x +private theorem List.append.eq_2.{u} : ∀ {α : Type u} (x : List α) (a : α) (l : List α), List.append (a :: l) x = a :: List.append l x -/ #guard_msgs in @@ -23,9 +23,9 @@ private theorem List.append._eq_2.{u} : ∀ {α : Type u} (x : List α) (a : α) /-- info: equations: -private theorem ack._eq_1 : ∀ (x : Nat), ack 0 x = x + 1 -private theorem ack._eq_2 : ∀ (x_2 : Nat), ack (Nat.succ x_2) 0 = ack x_2 1 -private theorem ack._eq_3 : ∀ (x_2 y : Nat), ack (Nat.succ x_2) (Nat.succ y) = ack x_2 (ack (x_2 + 1) y) +private theorem ack.eq_1 : ∀ (x : Nat), ack 0 x = x + 1 +private theorem ack.eq_2 : ∀ (x_2 : Nat), ack (Nat.succ x_2) 0 = ack x_2 1 +private theorem ack.eq_3 : ∀ (x_2 y : Nat), ack (Nat.succ x_2) (Nat.succ y) = ack x_2 (ack (x_2 + 1) y) -/ #guard_msgs in #print eqns ack diff --git a/tests/lean/run/robinson.lean b/tests/lean/run/robinson.lean index 2327bbc088fe..6f82a09e645c 100644 --- a/tests/lean/run/robinson.lean +++ b/tests/lean/run/robinson.lean @@ -49,10 +49,10 @@ decreasing_by attribute [simp] robinson set_option pp.proofs true -#check robinson._eq_1 -#check robinson._eq_2 -#check robinson._eq_3 -#check robinson._eq_4 +#check robinson.eq_1 +#check robinson.eq_2 +#check robinson.eq_3 +#check robinson.eq_4 theorem ex : (robinson (Term.Var 0) (Term.Var 0)).1 = some id := by unfold robinson diff --git a/tests/lean/run/rossel1.lean b/tests/lean/run/rossel1.lean index a62747d9568c..57837d4b8d7e 100644 --- a/tests/lean/run/rossel1.lean +++ b/tests/lean/run/rossel1.lean @@ -21,9 +21,9 @@ def Lineage.container' {rtr i} : Lineage rtr i → (Option ID × Reactor) attribute [simp] Lineage.container' -#check @Lineage.container'._eq_1 -#check @Lineage.container'._eq_2 -#check @Lineage.container'._eq_3 +#check @Lineage.container'.eq_1 +#check @Lineage.container'.eq_2 +#check @Lineage.container'.eq_3 @[simp] def Lineage.container : Lineage rtr i → (Option ID × Reactor) | nested l@h:(nested ..) _ => l.container diff --git a/tests/lean/run/simpAtDefIssue.lean b/tests/lean/run/simpAtDefIssue.lean index e82f81f901b9..bae33f330827 100644 --- a/tests/lean/run/simpAtDefIssue.lean +++ b/tests/lean/run/simpAtDefIssue.lean @@ -5,7 +5,7 @@ | x+1, 5 => 2 * g x 0 | x+1, y => 2 * g x y -#check g._eq_1 -#check g._eq_2 -#check g._eq_3 -#check g._eq_4 +#check g.eq_1 +#check g.eq_2 +#check g.eq_3 +#check g.eq_4 diff --git a/tests/lean/run/simp_eqn_bug.lean b/tests/lean/run/simp_eqn_bug.lean index 22f62770eb12..9871a9f8f5cf 100644 --- a/tests/lean/run/simp_eqn_bug.lean +++ b/tests/lean/run/simp_eqn_bug.lean @@ -5,6 +5,6 @@ | [] => a | x::xs => x + f a xs -example : f 25 xs = 0 := by apply f._eq_1 -example (h : a = 25 → False) : f a [] = a := by apply f._eq_2; assumption -example (h : a = 25 → False) : f a (x::xs) = x + f a xs := by apply f._eq_3; assumption +example : f 25 xs = 0 := by apply f.eq_1 +example (h : a = 25 → False) : f a [] = a := by apply f.eq_2; assumption +example (h : a = 25 → False) : f a (x::xs) = x + f a xs := by apply f.eq_3; assumption diff --git a/tests/lean/run/splitIssue.lean b/tests/lean/run/splitIssue.lean index fe8b4cbdc79d..143722e82909 100644 --- a/tests/lean/run/splitIssue.lean +++ b/tests/lean/run/splitIssue.lean @@ -19,9 +19,9 @@ theorem len_nil : len ([] : List α) = 0 := by simp [len] -- The `simp [len]` above generated the following equation theorems for len -#check @len._eq_1 -#check @len._eq_2 -#check @len._eq_3 -- It is conditional, and may be tricky to use. +#check @len.eq_1 +#check @len.eq_2 +#check @len.eq_3 -- It is conditional, and may be tricky to use. theorem len_1 (a : α) : len [a] = 1 := by simp [len] @@ -30,7 +30,7 @@ theorem len_2 (a b : α) (bs : List α) : len (a::b::bs) = 1 + len (b::bs) := by conv => lhs; unfold len -- The `unfold` tactic above generated the following theorem -#check @len._unfold +#check @len.def theorem len_cons (a : α) (as : List α) : len (a::as) = 1 + len as := by cases as with diff --git a/tests/lean/run/splitList.lean b/tests/lean/run/splitList.lean index 239a7916718d..240ec9569812 100644 --- a/tests/lean/run/splitList.lean +++ b/tests/lean/run/splitList.lean @@ -38,9 +38,9 @@ theorem len_nil : len ([] : List α) = 0 := by simp [len] -- The `simp [len]` above generated the following equation theorems for len -#check @len._eq_1 -#check @len._eq_2 -#check @len._eq_3 +#check @len.eq_1 +#check @len.eq_2 +#check @len.eq_3 theorem len_1 (a : α) : len [a] = 1 := by simp [len] @@ -49,7 +49,7 @@ theorem len_2 (a b : α) (bs : List α) : len (a::b::bs) = 1 + len (b::bs) := by conv => lhs; unfold len -- The `unfold` tactic above generated the following theorem -#check @len._unfold +#check @len.def theorem len_cons (a : α) (as : List α) : len (a::as) = 1 + len as := by cases as with @@ -88,9 +88,9 @@ theorem len_nil : len ([] : List α) = 0 := by simp [len] -- The `simp [len]` above generated the following equation theorems for len -#check @len._eq_1 -#check @len._eq_2 -#check @len._eq_3 +#check @len.eq_1 +#check @len.eq_2 +#check @len.eq_3 theorem len_1 (a : α) : len [a] = 1 := by simp [len] @@ -99,7 +99,7 @@ theorem len_2 (a b : α) (bs : List α) : len (a::b::bs) = 1 + len (b::bs) := by conv => lhs; unfold len -- The `unfold` tactic above generated the following theorem -#check @len._unfold +#check @len.def theorem len_cons (a : α) (as : List α) : len (a::as) = 1 + len as := by cases as with diff --git a/tests/lean/run/streamEqIssue.lean b/tests/lean/run/streamEqIssue.lean index 6401c0016ebe..2174a73bf3ce 100644 --- a/tests/lean/run/streamEqIssue.lean +++ b/tests/lean/run/streamEqIssue.lean @@ -4,6 +4,6 @@ | n + 1, some (_, s') => hasLength n s' | _, _ => false -#check @Stream.hasLength._eq_1 -#check @Stream.hasLength._eq_2 -#check @Stream.hasLength._eq_3 +#check @Stream.hasLength.eq_1 +#check @Stream.hasLength.eq_2 +#check @Stream.hasLength.eq_3 diff --git a/tests/lean/structuralEqns.lean b/tests/lean/run/structEqns.lean similarity index 55% rename from tests/lean/structuralEqns.lean rename to tests/lean/run/structEqns.lean index 4bcb423ec99e..f5842c62f828 100644 --- a/tests/lean/structuralEqns.lean +++ b/tests/lean/run/structEqns.lean @@ -15,11 +15,32 @@ def foo (xs ys zs : List Nat) : List Nat := | _ => [2] #eval tst ``foo -#check foo._unfold + +/-- +info: foo.def (xs ys zs : List Nat) : + foo xs ys zs = + match (xs, ys) with + | (xs', ys') => + match zs with + | z :: zs => foo xs ys zs + | x => + match ys' with + | [] => [1] + | x => [2] +-/ +#guard_msgs in +#check foo.def def bar (xs ys : List Nat) : List Nat := match xs ++ [], ys ++ [] with | xs', ys' => xs' ++ ys' +/-- +info: def bar : List Nat → List Nat → List Nat := +fun xs ys => + match xs ++ [], ys ++ [] with + | xs', ys' => xs' ++ ys' +-/ +#guard_msgs in #print bar -- should not contain either `let _discr` aux binding diff --git a/tests/lean/run/structuralEqns.lean b/tests/lean/run/structuralEqns.lean index f9e1892554dd..c3ace8752f54 100644 --- a/tests/lean/run/structuralEqns.lean +++ b/tests/lean/run/structuralEqns.lean @@ -6,9 +6,9 @@ def tst (declName : Name) : MetaM Unit := do IO.println (← getUnfoldEqnFor? declName) #eval tst ``List.map -#check @List.map._eq_1 -#check @List.map._eq_2 -#check @List.map._unfold +#check @List.map.eq_1 +#check @List.map.eq_2 +#check @List.map.def def foo (xs ys zs : List Nat) : List Nat := match (xs, ys) with @@ -21,9 +21,9 @@ def foo (xs ys zs : List Nat) : List Nat := #eval tst ``foo -#check foo._eq_1 -#check foo._eq_2 -#check foo._unfold +#check foo.eq_1 +#check foo.eq_2 +#check foo.def #eval tst ``foo @@ -35,12 +35,12 @@ def g : List Nat → List Nat → Nat | x::xs, [] => g xs [] #eval tst ``g -#check g._eq_1 -#check g._eq_2 -#check g._eq_3 -#check g._eq_4 -#check g._eq_5 -#check g._unfold +#check g.eq_1 +#check g.eq_2 +#check g.eq_3 +#check g.eq_4 +#check g.eq_5 +#check g.def def h (xs : List Nat) (y : Nat) : Nat := match xs with @@ -51,9 +51,9 @@ def h (xs : List Nat) (y : Nat) : Nat := | y+1 => h xs y #eval tst ``h -#check h._eq_1 -#check h._eq_2 -#check h._unfold +#check h.eq_1 +#check h.eq_2 +#check h.def def r (i j : Nat) : Nat := i + @@ -65,10 +65,10 @@ def r (i j : Nat) : Nat := | Nat.succ j => r i j #eval tst ``r -#check r._eq_1 -#check r._eq_2 -#check r._eq_3 -#check r._unfold +#check r.eq_1 +#check r.eq_2 +#check r.eq_3 +#check r.def def bla (f g : α → α → α) (a : α) (i : α) (j : Nat) : α := f i <| @@ -80,7 +80,7 @@ def bla (f g : α → α → α) (a : α) (i : α) (j : Nat) : α := | Nat.succ j => bla f g a i j #eval tst ``bla -#check @bla._eq_1 -#check @bla._eq_2 -#check @bla._eq_3 -#check @bla._unfold +#check @bla.eq_1 +#check @bla.eq_2 +#check @bla.eq_3 +#check @bla.def diff --git a/tests/lean/run/structuralEqns2.lean b/tests/lean/run/structuralEqns2.lean index 4d9d9a1fac41..83f8dd379192 100644 --- a/tests/lean/run/structuralEqns2.lean +++ b/tests/lean/run/structuralEqns2.lean @@ -12,9 +12,9 @@ def g (i j : Nat) : Nat := | Nat.succ j => g i j #eval tst ``g -#check g._eq_1 -#check g._eq_2 -#check g._unfold +#check g.eq_1 +#check g.eq_2 +#check g.def def h (i j : Nat) : Nat := let z := @@ -24,6 +24,6 @@ def h (i j : Nat) : Nat := z + z #eval tst ``h -#check h._eq_1 -#check h._eq_2 -#check h._unfold +#check h.eq_1 +#check h.eq_2 +#check h.def diff --git a/tests/lean/run/structuralEqns3.lean b/tests/lean/run/structuralEqns3.lean index ac72ee052221..09aa049b215c 100644 --- a/tests/lean/run/structuralEqns3.lean +++ b/tests/lean/run/structuralEqns3.lean @@ -15,6 +15,6 @@ def wk_comp : Wk n m → Wk m l → Wk n l #eval tst ``wk_comp -#check @wk_comp._eq_1 -#check @wk_comp._eq_2 -#check @wk_comp._unfold +#check @wk_comp.eq_1 +#check @wk_comp.eq_2 +#check @wk_comp.def diff --git a/tests/lean/run/wfEqns1.lean b/tests/lean/run/wfEqns1.lean index 1e3e4fb7684a..598d1692f4f2 100644 --- a/tests/lean/run/wfEqns1.lean +++ b/tests/lean/run/wfEqns1.lean @@ -22,6 +22,6 @@ end #print isEven #eval tst ``isEven -#check @isEven._eq_1 -#check @isEven._eq_2 -#check @isEven._unfold +#check @isEven.eq_1 +#check @isEven.eq_2 +#check @isEven.def diff --git a/tests/lean/run/wfEqns2.lean b/tests/lean/run/wfEqns2.lean index aaf8670cb3aa..c0b5dfc117c3 100644 --- a/tests/lean/run/wfEqns2.lean +++ b/tests/lean/run/wfEqns2.lean @@ -31,10 +31,10 @@ decreasing_by end #eval tst ``g -#check g._eq_1 -#check g._eq_2 -#check g._unfold +#check g.eq_1 +#check g.eq_2 +#check g.def #eval tst ``h -#check h._eq_1 -#check h._eq_2 -#check h._unfold +#check h.eq_1 +#check h.eq_2 +#check h.def diff --git a/tests/lean/run/wfEqns3.lean b/tests/lean/run/wfEqns3.lean index 850766036ead..b57d06aa4cab 100644 --- a/tests/lean/run/wfEqns3.lean +++ b/tests/lean/run/wfEqns3.lean @@ -15,5 +15,5 @@ decreasing_by exact h #eval tst ``f -#check f._eq_1 -#check f._unfold +#check f.eq_1 +#check f.def diff --git a/tests/lean/run/wfEqns4.lean b/tests/lean/run/wfEqns4.lean index d6549693aaa2..7b683f879e6e 100644 --- a/tests/lean/run/wfEqns4.lean +++ b/tests/lean/run/wfEqns4.lean @@ -37,12 +37,12 @@ end #eval f 5 'a' 'b' #eval tst ``f -#check @f._eq_1 -#check @f._eq_2 -#check @f._unfold +#check @f.eq_1 +#check @f.eq_2 +#check @f.def #eval tst ``h -#check @h._eq_1 -#check @h._eq_2 -#check @h._unfold +#check @h.eq_1 +#check @h.eq_2 +#check @h.def diff --git a/tests/lean/structuralEqns.lean.expected.out b/tests/lean/structuralEqns.lean.expected.out deleted file mode 100644 index 5e1dc23bdfdd..000000000000 --- a/tests/lean/structuralEqns.lean.expected.out +++ /dev/null @@ -1,15 +0,0 @@ -(some _private.structuralEqns.0.foo._unfold) -foo._unfold (xs ys zs : List Nat) : - foo xs ys zs = - match (xs, ys) with - | (xs', ys') => - match zs with - | z :: zs => foo xs ys zs - | x => - match ys' with - | [] => [1] - | x => [2] -def bar : List Nat → List Nat → List Nat := -fun xs ys => - match xs ++ [], ys ++ [] with - | xs', ys' => xs' ++ ys' From a81205c290df38d87d9ac9ed70cb875038475253 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Wed, 13 Mar 2024 10:03:39 +0100 Subject: [PATCH 21/32] feat: conv => calc (#3659) `calc` is great for explicit rewriting, `conv` is great to say where to rewrite, so it's natural to want `calc` as a `conv` tactic. Zulip disucssion at https://leanprover.zulipchat.com/#narrow/stream/270676-lean4/topic/idiom.20for.20using.20calc.20to.20rewrite.20the.20goal/near/424269608 Fixes #3557. --- src/Init/Conv.lean | 1 - src/Init/NotationExtra.lean | 45 ++++++------------------------------ tests/lean/run/convcalc.lean | 6 +++++ 3 files changed, 13 insertions(+), 39 deletions(-) create mode 100644 tests/lean/run/convcalc.lean diff --git a/src/Init/Conv.lean b/src/Init/Conv.lean index ee45526b29a9..5eccbdcb5ff7 100644 --- a/src/Init/Conv.lean +++ b/src/Init/Conv.lean @@ -156,7 +156,6 @@ match [a, b] with simplifies to `a`. -/ syntax (name := simpMatch) "simp_match" : conv - /-- Executes the given tactic block without converting `conv` goal into a regular goal. -/ syntax (name := nestedTacticCore) "tactic'" " => " tacticSeq : conv diff --git a/src/Init/NotationExtra.lean b/src/Init/NotationExtra.lean index e728e723fc70..c193eaf9585c 100644 --- a/src/Init/NotationExtra.lean +++ b/src/Init/NotationExtra.lean @@ -9,6 +9,7 @@ prelude import Init.Meta import Init.Data.Array.Subarray import Init.Data.ToString +import Init.Conv namespace Lean macro "Macro.trace[" id:ident "]" s:interpolatedStr(term) : term => @@ -123,7 +124,7 @@ calc abc _ = xyz := pwxyz ``` -`calc` has term mode and tactic mode variants. This is the term mode variant. +`calc` works as a term, as a tactic or as a `conv` tactic. See [Theorem Proving in Lean 4][tpil4] for more information. @@ -131,45 +132,13 @@ See [Theorem Proving in Lean 4][tpil4] for more information. -/ syntax (name := calc) "calc" calcSteps : term -/-- Step-wise reasoning over transitive relations. -``` -calc - a = b := pab - b = c := pbc - ... - y = z := pyz -``` -proves `a = z` from the given step-wise proofs. `=` can be replaced with any -relation implementing the typeclass `Trans`. Instead of repeating the right- -hand sides, subsequent left-hand sides can be replaced with `_`. -``` -calc - a = b := pab - _ = c := pbc - ... - _ = z := pyz -``` -It is also possible to write the *first* relation as `\n _ = := -`. This is useful for aligning relation symbols: -``` -calc abc - _ = bce := pabce - _ = cef := pbcef - ... - _ = xyz := pwxyz -``` - -`calc` has term mode and tactic mode variants. This is the tactic mode variant, -which supports an additional feature: it works even if the goal is `a = z'` -for some other `z'`; in this case it will not close the goal but will instead -leave a subgoal proving `z = z'`. - -See [Theorem Proving in Lean 4][tpil4] for more information. - -[tpil4]: https://lean-lang.org/theorem_proving_in_lean4/quantifiers_and_equality.html#calculational-proofs --/ +@[inherit_doc «calc»] syntax (name := calcTactic) "calc" calcSteps : tactic +@[inherit_doc «calc»] +macro tk:"calc" steps:calcSteps : conv => + `(conv| tactic => calc%$tk $steps) + @[app_unexpander Unit.unit] def unexpandUnit : Lean.PrettyPrinter.Unexpander | `($(_)) => `(()) diff --git a/tests/lean/run/convcalc.lean b/tests/lean/run/convcalc.lean new file mode 100644 index 000000000000..f21348d4c0d5 --- /dev/null +++ b/tests/lean/run/convcalc.lean @@ -0,0 +1,6 @@ +example (P : Nat → Prop) (h : P 4) : P ((1 + 1) + 2) := by + conv => arg 1; calc + (1 + 1) + 2 + _ = 2 + 2 := by simp + _ = 4 := by simp + exact h From 600412838c039c942cc92a4332fc8cbe63f2f309 Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Wed, 13 Mar 2024 02:38:37 -0700 Subject: [PATCH 22/32] fix: auxiliary definition nested in theorem should be `def` if its type is not a proposition (#3662) --- src/Lean/Elab/MutualDef.lean | 6 +++++- tests/lean/run/letrecInThm.lean | 23 +++++++++++++++++++++++ 2 files changed, 28 insertions(+), 1 deletion(-) create mode 100644 tests/lean/run/letrecInThm.lean diff --git a/src/Lean/Elab/MutualDef.lean b/src/Lean/Elab/MutualDef.lean index d58193aa5eb7..2413acfa3a87 100644 --- a/src/Lean/Elab/MutualDef.lean +++ b/src/Lean/Elab/MutualDef.lean @@ -657,10 +657,14 @@ def pushLetRecs (preDefs : Array PreDefinition) (letRecClosures : List LetRecClo letRecClosures.foldlM (init := preDefs) fun preDefs c => do let type := Closure.mkForall c.localDecls c.toLift.type let value := Closure.mkLambda c.localDecls c.toLift.val - -- Convert any proof let recs inside a `def` to `theorem` kind let kind ← if kind.isDefOrAbbrevOrOpaque then + -- Convert any proof let recs inside a `def` to `theorem` kind withLCtx c.toLift.lctx c.toLift.localInstances do return if (← inferType c.toLift.type).isProp then .theorem else kind + else if kind.isTheorem then + -- Convert any non-proof let recs inside a `theorem` to `def` kind + withLCtx c.toLift.lctx c.toLift.localInstances do + return if (← inferType c.toLift.type).isProp then .theorem else .def else pure kind return preDefs.push { diff --git a/tests/lean/run/letrecInThm.lean b/tests/lean/run/letrecInThm.lean new file mode 100644 index 000000000000..8fad5b8767e0 --- /dev/null +++ b/tests/lean/run/letrecInThm.lean @@ -0,0 +1,23 @@ +-- Auxiliary definitions nested in theorems must be defs + +theorem foo : 10 = 10 := rfl +where aux : Nat := 20 + +/-- +info: def foo.aux : Nat := +20 +-/ +#guard_msgs in +#print foo.aux + + +theorem foo2 : 10 = 10 := + let rec aux (x : Nat) : Nat := x + 1 + rfl + +/-- +info: def foo2.aux : Nat → Nat := +fun x => x + 1 +-/ +#guard_msgs in +#print foo2.aux From e61d082a9555464bca6de954f8fbc2a1b1a1bf89 Mon Sep 17 00:00:00 2001 From: Hongyu Ouyang <96765450+casavaca@users.noreply.github.com> Date: Wed, 13 Mar 2024 03:51:24 -0700 Subject: [PATCH 23/32] doc: fix typo in USize.size docstring (#3664) --- src/Init/Prelude.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Init/Prelude.lean b/src/Init/Prelude.lean index cd7f37c01f1c..918ff44428bc 100644 --- a/src/Init/Prelude.lean +++ b/src/Init/Prelude.lean @@ -2036,7 +2036,7 @@ instance : Inhabited UInt64 where default := UInt64.ofNatCore 0 (by decide) /-- -The size of type `UInt16`, that is, `2^System.Platform.numBits`, which may +The size of type `USize`, that is, `2^System.Platform.numBits`, which may be either `2^32` or `2^64` depending on the platform's architecture. Remark: we define `USize.size` using `(2^numBits - 1) + 1` to ensure the From 84b0919a116e9be12f933e764474f45d964ce85c Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Tue, 12 Mar 2024 18:22:31 -0700 Subject: [PATCH 24/32] feat: type of theorems must be propositions --- RELEASES.md | 2 + src/Lean/Elab/MutualDef.lean | 3 + src/Lean/Environment.lean | 1 + src/Lean/Message.lean | 1 + src/kernel/environment.cpp | 2 + src/kernel/kernel_exception.h | 33 +++++++---- tests/lean/caseSuggestions.lean | 12 ++-- tests/lean/levelNGen.lean | 6 -- tests/lean/levelNGen.lean.expected.out | 5 -- tests/lean/run/contra.lean | 30 +++++----- tests/lean/run/discrRefinement2.lean | 4 +- tests/lean/run/inj2.lean | 2 +- tests/lean/run/lazyListRotateUnfoldProof.lean | 2 +- tests/lean/run/lemma.lean | 12 +++- tests/lean/run/levelNGen.lean | 14 +++++ tests/lean/run/newfrontend3.lean | 2 +- tests/lean/run/structuralRec1.lean | 26 ++++----- tests/lean/run/tactic1.lean | 4 +- tests/lean/run/thmIsProp.lean | 57 +++++++++++++++++++ 19 files changed, 153 insertions(+), 65 deletions(-) delete mode 100644 tests/lean/levelNGen.lean delete mode 100644 tests/lean/levelNGen.lean.expected.out create mode 100644 tests/lean/run/levelNGen.lean create mode 100644 tests/lean/run/thmIsProp.lean diff --git a/RELEASES.md b/RELEASES.md index e6c7e073ddc0..f23616684e54 100644 --- a/RELEASES.md +++ b/RELEASES.md @@ -11,6 +11,8 @@ of each version. v4.8.0 (development in progress) --------- +* Lean now generates an error if the type of a theorem is **not** a proposition. + * New command `derive_functinal_induction`: Derived from the definition of a (possibly mutually) recursive function diff --git a/src/Lean/Elab/MutualDef.lean b/src/Lean/Elab/MutualDef.lean index 2413acfa3a87..564d0de9608f 100644 --- a/src/Lean/Elab/MutualDef.lean +++ b/src/Lean/Elab/MutualDef.lean @@ -644,6 +644,9 @@ def pushMain (preDefs : Array PreDefinition) (sectionVars : Array Expr) (mainHea let termination := termination.rememberExtraParams header.numParams mainVals[i]! let value ← mkLambdaFVars sectionVars mainVals[i]! let type ← mkForallFVars sectionVars header.type + if header.kind.isTheorem then + unless (← isProp type) do + throwErrorAt header.ref "type of theorem '{header.declName}' is not a proposition{indentExpr type}" return preDefs.push { ref := getDeclarationSelectionRef header.ref kind := header.kind diff --git a/src/Lean/Environment.lean b/src/Lean/Environment.lean index 604c0ab30569..8c5f6dd6ad23 100644 --- a/src/Lean/Environment.lean +++ b/src/Lean/Environment.lean @@ -230,6 +230,7 @@ inductive KernelException where | exprTypeMismatch (env : Environment) (lctx : LocalContext) (expr : Expr) (expectedType : Expr) | appTypeMismatch (env : Environment) (lctx : LocalContext) (app : Expr) (funType : Expr) (argType : Expr) | invalidProj (env : Environment) (lctx : LocalContext) (proj : Expr) + | thmTypeIsNotProp (env : Environment) (name : Name) (type : Expr) | other (msg : String) | deterministicTimeout | excessiveMemory diff --git a/src/Lean/Message.lean b/src/Lean/Message.lean index d51a1d9c5d94..76d1901b131b 100644 --- a/src/Lean/Message.lean +++ b/src/Lean/Message.lean @@ -366,6 +366,7 @@ def toMessageData (e : KernelException) (opts : Options) : MessageData := | appTypeMismatch env lctx e fnType argType => mkCtx env lctx opts m!"application type mismatch{indentExpr e}\nargument has type{indentExpr argType}\nbut function has type{indentExpr fnType}" | invalidProj env lctx e => mkCtx env lctx opts m!"(kernel) invalid projection{indentExpr e}" + | thmTypeIsNotProp env constName type => mkCtx env {} opts m!"(kernel) type of theorem '{constName}' is not a proposition{indentExpr type}" | other msg => m!"(kernel) {msg}" | deterministicTimeout => "(kernel) deterministic timeout" | excessiveMemory => "(kernel) excessive memory consumption detected" diff --git a/src/kernel/environment.cpp b/src/kernel/environment.cpp index 56b67af4a7c2..792f4880dfe3 100644 --- a/src/kernel/environment.cpp +++ b/src/kernel/environment.cpp @@ -178,6 +178,8 @@ environment environment::add_theorem(declaration const & d, bool check) const { if (check) { // TODO(Leo): we must add support for handling tasks here type_checker checker(*this); + if (!checker.is_prop(v.get_type())) + throw theorem_type_is_not_prop(*this, v.get_name(), v.get_type()); check_constant_val(*this, v.to_constant_val(), checker); check_no_metavar_no_fvar(*this, v.get_name(), v.get_value()); expr val_type = checker.check(v.get_value(), v.get_lparams()); diff --git a/src/kernel/kernel_exception.h b/src/kernel/kernel_exception.h index 930185828fbd..793cf6403277 100644 --- a/src/kernel/kernel_exception.h +++ b/src/kernel/kernel_exception.h @@ -66,6 +66,16 @@ class declaration_has_free_vars_exception : public kernel_exception { expr const & get_expr() const { return m_expr; } }; +class theorem_type_is_not_prop : public kernel_exception { + name m_name; + expr m_type; +public: + theorem_type_is_not_prop(environment const & env, name const & n, expr const & type): + kernel_exception(env), m_name(n), m_type(type) {} + name const & get_decl_name() const { return m_name; } + expr const & get_type() const { return m_type; } +}; + class kernel_exception_with_lctx : public kernel_exception { local_ctx m_lctx; public: @@ -185,21 +195,24 @@ object * catch_kernel_exceptions(std::function const & f) { } catch (invalid_proj_exception & ex) { // 10 | invalidProj (env : Environment) (lctx : LocalContext) (proj : Expr) return mk_cnstr(0, mk_cnstr(10, ex.env(), ex.get_local_ctx(), ex.get_proj())).steal(); + } catch (theorem_type_is_not_prop & ex) { + // 11 | thmTypeIsNotProp (env : Environment) (name : Name) (type : Expr) + return mk_cnstr(0, mk_cnstr(11, ex.env(), ex.get_decl_name(), ex.get_type())).steal(); } catch (exception & ex) { - // 11 | other (msg : String) - return mk_cnstr(0, mk_cnstr(11, string_ref(ex.what()))).steal(); + // 12 | other (msg : String) + return mk_cnstr(0, mk_cnstr(12, string_ref(ex.what()))).steal(); } catch (heartbeat_exception & ex) { - // 12 | deterministicTimeout - return mk_cnstr(0, box(12)).steal(); - } catch (memory_exception & ex) { - // 13 | excessiveMemory + // 13 | deterministicTimeout return mk_cnstr(0, box(13)).steal(); - } catch (stack_space_exception & ex) { - // 14 | deepRecursion + } catch (memory_exception & ex) { + // 14 | excessiveMemory return mk_cnstr(0, box(14)).steal(); - } catch (interrupted & ex) { - // 15 | interrupted + } catch (stack_space_exception & ex) { + // 15 | deepRecursion return mk_cnstr(0, box(15)).steal(); + } catch (interrupted & ex) { + // 16 | interrupted + return mk_cnstr(0, box(16)).steal(); } } } diff --git a/tests/lean/caseSuggestions.lean b/tests/lean/caseSuggestions.lean index 5b212658cc5a..fca6d7c5bec6 100644 --- a/tests/lean/caseSuggestions.lean +++ b/tests/lean/caseSuggestions.lean @@ -7,7 +7,7 @@ /-! This example tests what happens when no cases are available. -/ -theorem noCases : Nat := by +def noCases : Nat := by case nonexistent => skip @@ -15,7 +15,7 @@ theorem noCases : Nat := by This example tests what happens when just one case is available, but it wasn't picked. -/ -theorem oneCase : Nat := by +def oneCase : Nat := by cases () case nonexistent => skip @@ -24,22 +24,22 @@ theorem oneCase : Nat := by Check varying numbers of cases to make sure the pretty-print setup for the list is correct. -/ -theorem twoCases : Nat := by +def twoCases : Nat := by cases true case nonexistent => skip -theorem fourCases : Nat := by +def fourCases : Nat := by cases true <;> cases true case nonexistent => skip -theorem eightCases : Nat := by +def eightCases : Nat := by cases true <;> cases true <;> cases true case nonexistent => skip -theorem sixteenCases : Nat := by +def sixteenCases : Nat := by cases true <;> cases true <;> cases true <;> cases true case nonexistent => skip diff --git a/tests/lean/levelNGen.lean b/tests/lean/levelNGen.lean deleted file mode 100644 index c071d396f088..000000000000 --- a/tests/lean/levelNGen.lean +++ /dev/null @@ -1,6 +0,0 @@ -opaque test1 {α : Sort _} : α → Sort u_1 -#check test1 -def test2 {α : Sort _} : α → Sort u_1 := sorry -#check test2 -variable {α : Sort _} in theorem test3 : α → Sort _ := sorry -#check test3 diff --git a/tests/lean/levelNGen.lean.expected.out b/tests/lean/levelNGen.lean.expected.out deleted file mode 100644 index 927f2ce74564..000000000000 --- a/tests/lean/levelNGen.lean.expected.out +++ /dev/null @@ -1,5 +0,0 @@ -test1.{u_1, u_2} {α : Sort u_2} : α → Sort u_1 -levelNGen.lean:3:4-3:9: warning: declaration uses 'sorry' -test2.{u_1, u_2} {α : Sort u_2} : α → Sort u_1 -levelNGen.lean:5:33-5:38: warning: declaration uses 'sorry' -test3.{u_1, u_2} {α : Sort u_2} : α → Sort u_1 diff --git a/tests/lean/run/contra.lean b/tests/lean/run/contra.lean index 5ffdf894f7bc..68c8f5dfdf62 100644 --- a/tests/lean/run/contra.lean +++ b/tests/lean/run/contra.lean @@ -1,46 +1,46 @@ -theorem ex1 (p : Prop) (h1 : p) (h2 : p → False) : α := by +def ex1 (p : Prop) (h1 : p) (h2 : p → False) : α := by contradiction -theorem ex2 (p : Prop) (h1 : p) (h2 : ¬ p) : α := by +def ex2 (p : Prop) (h1 : p) (h2 : ¬ p) : α := by contradiction -theorem ex3 (p : Prop) (h1 : id p) (h2 : ¬ p) : α := by +def ex3 (p : Prop) (h1 : id p) (h2 : ¬ p) : α := by contradiction -theorem ex4 (p : Prop) (h1 : id p) (h2 : id (Not p)) : α := by +def ex4 (p : Prop) (h1 : id p) (h2 : id (Not p)) : α := by contradiction -theorem ex5 (h : x+1 = 0) : α := by +def ex5 (h : x+1 = 0) : α := by contradiction -theorem ex6 (h : 0+0 ≠ 0) : α := by +def ex6 (h : 0+0 ≠ 0) : α := by contradiction -theorem ex7 (x : α) (h : Not (x = x)) : α := by +def ex7 (x : α) (h : Not (x = x)) : α := by contradiction -theorem ex8 (h : 0+0 = 0 → False) : α := by +def ex8 (h : 0+0 = 0 → False) : α := by contradiction -theorem ex9 (h : 10 = 20) : α := by +def ex9 (h : 10 = 20) : α := by contradiction -theorem ex10 (h : [] = [1, 2, 3]) : α := by +def ex10 (h : [] = [1, 2, 3]) : α := by contradiction -theorem ex11 (h : id [] = [1, 2, 3]) : α := by +def ex11 (h : id [] = [1, 2, 3]) : α := by contradiction -theorem ex12 (h : False) : α := by +def ex12 (h : False) : α := by contradiction -theorem ex13 (h : id False) : α := by +def ex13 (h : id False) : α := by contradiction -theorem ex14 (h : 100000000 ≤ 20) : α := by +def ex14 (h : 100000000 ≤ 20) : α := by contradiction -theorem ex15 (x : α) (h : x = x → False) : α := by +def ex15 (x : α) (h : x = x → False) : α := by contradiction theorem ex16 (xs : List α) (h : xs = [] → False) : Nonempty α := by diff --git a/tests/lean/run/discrRefinement2.lean b/tests/lean/run/discrRefinement2.lean index 56ee80495d88..ccaef6d8621f 100644 --- a/tests/lean/run/discrRefinement2.lean +++ b/tests/lean/run/discrRefinement2.lean @@ -3,11 +3,11 @@ theorem ex1 {α : Sort u} {a b : α} (h : a ≅ b) : a = b := match h with | HEq.refl _ => rfl -theorem ex2 {α : Sort u2} {a : α} {motive : {β : Sort u2} → β → Sort u1} (m : motive a) {β : Sort u2} {b : β} (h : a ≅ b) : motive b := +def ex2 {α : Sort u2} {a : α} {motive : {β : Sort u2} → β → Sort u1} (m : motive a) {β : Sort u2} {b : β} (h : a ≅ b) : motive b := match h, m with | HEq.refl _, m => m -theorem ex3 {α : Sort u} {a : α} {p : α → Sort v} {b : α} (h₁ : a ≅ b) (h₂ : p a) : p b := +def ex3 {α : Sort u} {a : α} {p : α → Sort v} {b : α} (h₁ : a ≅ b) (h₂ : p a) : p b := match h₁, h₂ with | HEq.refl _, h₂ => h₂ diff --git a/tests/lean/run/inj2.lean b/tests/lean/run/inj2.lean index 0c4e57d3bd93..a129b8386edd 100644 --- a/tests/lean/run/inj2.lean +++ b/tests/lean/run/inj2.lean @@ -24,7 +24,7 @@ by { rw [h4] } -theorem test4 {α} (v : Fin2 0) : α := +def test4 {α} (v : Fin2 0) : α := by cases v def test5 {α β} {n} (v : Vec2 α β (n+1)) : α := by diff --git a/tests/lean/run/lazyListRotateUnfoldProof.lean b/tests/lean/run/lazyListRotateUnfoldProof.lean index 517181cb66c7..4946077ec40a 100644 --- a/tests/lean/run/lazyListRotateUnfoldProof.lean +++ b/tests/lean/run/lazyListRotateUnfoldProof.lean @@ -30,7 +30,7 @@ theorem rotate_inv {F : LazyList τ} {R : List τ} : (h : F.length + 1 = R.lengt | LazyList.cons Fh Ft => sorry | LazyList.delayed Ft => sorry -theorem LazyList.ind {α : Type u} {motive : LazyList α → Sort v} +def LazyList.ind {α : Type u} {motive : LazyList α → Sort v} (nil : motive LazyList.nil) (cons : (hd : α) → (tl : LazyList α) → motive tl → motive (LazyList.cons hd tl)) (delayed : (t : Thunk (LazyList α)) → motive t.get → motive (LazyList.delayed t)) diff --git a/tests/lean/run/lemma.lean b/tests/lean/run/lemma.lean index 31fffea0bad1..e6739dd6a67c 100644 --- a/tests/lean/run/lemma.lean +++ b/tests/lean/run/lemma.lean @@ -1,8 +1,14 @@ macro mods:declModifiers "lemma" n:declId sig:declSig val:declVal : command => `($mods:declModifiers theorem $n $sig $val) -lemma fooSimple (n : Nat) : Prop := +def fooSimple (n : Nat) : Prop := if n = 0 then True else False -lemma fooPat : Nat → Prop +lemma fooSimple' : fooSimple 0 := + by constructor + +def fooPat : Nat → Prop | 0 => True - | n+1 => False + | _+1 => False + +lemma fooPat' : (x : Nat) → fooPat x → x = 0 + | 0, _ => rfl diff --git a/tests/lean/run/levelNGen.lean b/tests/lean/run/levelNGen.lean new file mode 100644 index 000000000000..5169b221c538 --- /dev/null +++ b/tests/lean/run/levelNGen.lean @@ -0,0 +1,14 @@ +opaque test1 {α : Sort _} : α → Sort u_1 +/-- info: test1.{u_1, u_2} {α : Sort u_2} : α → Sort u_1 -/ +#guard_msgs in +#check test1 + +def test2 {α : Sort _} : α → Sort u_1 := sorry +/-- info: test2.{u_1, u_2} {α : Sort u_2} : α → Sort u_1 -/ +#guard_msgs in +#check test2 + +variable {α : Sort _} in def test3 : α → Sort _ := sorry +/-- info: test3.{u_1, u_2} {α : Sort u_1} : α → Sort u_2 -/ +#guard_msgs in +#check test3 diff --git a/tests/lean/run/newfrontend3.lean b/tests/lean/run/newfrontend3.lean index d2a80068b3ff..be3efd0cc7df 100644 --- a/tests/lean/run/newfrontend3.lean +++ b/tests/lean/run/newfrontend3.lean @@ -9,7 +9,7 @@ def f (h : Nat → ({α : Type} → α → α) × Bool) : Nat := def tst : Nat := f fun n => (fun x => x, true) -theorem ex : id (Nat → Nat) := +def ex : id (Nat → Nat) := by { intro; assumption diff --git a/tests/lean/run/structuralRec1.lean b/tests/lean/run/structuralRec1.lean index b231217a5b31..8015f860dc0d 100644 --- a/tests/lean/run/structuralRec1.lean +++ b/tests/lean/run/structuralRec1.lean @@ -35,7 +35,7 @@ loop as def pmap2 {α β} (f : α → β) (as : PList α) : PList β := let rec loop : PList α → PList β - | PList.nil => PList.nil + | PList.nil => PList.nil | a:::as => f a ::: loop as; loop as @@ -58,7 +58,7 @@ match xs with | x:::xs => let y := 2 * x; match xs with - | PList.nil => PList.nil + | PList.nil => PList.nil | x:::xs => (y + x) ::: pfoo xs #eval foo [1, 2, 3, 4] @@ -79,11 +79,11 @@ else def pbla (x : Nat) (ys : PList Nat) : PList Nat := if x % 2 == 0 then match ys with - | PList.nil => PList.nil + | PList.nil => PList.nil | y:::ys => (y + x/2) ::: pbla (x/2) ys else match ys with - | PList.nil => PList.nil + | PList.nil => PList.nil | y:::ys => (y + x/2 + 1) ::: pbla (x/2) ys theorem blaEq (y : Nat) (ys : List Nat) : bla 4 (y::ys) = (y+2) :: bla 2 ys := @@ -121,7 +121,7 @@ def pg (xs : PList Nat) : True := | y:::ys => match ys with | PList.nil => True.intro - | _ => pg ys + | _ => pg ys def aux : Nat → Nat → Nat | 0, y => y @@ -157,7 +157,7 @@ axiom F0 : P 0 axiom F1 : P (F 0) axiom FS {n : Nat} : P n → P (F (F n)) -axiom T : Nat → Type +axiom T : Nat → Prop axiom TF0 : T 0 axiom TF1 : T (F 0) axiom TFS {n : Nat} : T n → T (F (F n)) @@ -175,13 +175,13 @@ theorem «nested recursion» : ∀ {n}, is_nat n → P n -- | _, is_nat.S .(is_nat.Z) => F1 -- | _, is_nat.S (is_nat.S h) => FS («nested recursion, inaccessible» h) -theorem «reordered discriminants, type» : ∀ n, is_nat_T n → Nat → T n := fun n hn m => +theorem «reordered discriminants, type» : ∀ n, is_nat_T n → Nat → T n := fun n hn m => match n, m, hn with | _, _, is_nat_T.Z => TF0 | _, _, is_nat_T.S is_nat_T.Z => TF1 | _, m, is_nat_T.S (is_nat_T.S h) => TFS («reordered discriminants, type» _ h m) -theorem «reordered discriminants» : ∀ n, is_nat n → Nat → P n := fun n hn m => +theorem «reordered discriminants» : ∀ n, is_nat n → Nat → P n := fun n hn m => match n, m, hn with | _, _, is_nat.Z => F0 | _, _, is_nat.S is_nat.Z => F1 @@ -194,8 +194,8 @@ match n, m, hn with -- | y::ys => -- match ys with -- | List.nil => True.intro --- | _::_::zs => «unsupported nesting» zs --- | zs => «unsupported nesting» ys +-- | _::_::zs => «unsupported nesting» zs +-- | zs => «unsupported nesting» ys def «unsupported nesting, predicate» (xs : PList Nat) : True := match xs with @@ -203,8 +203,8 @@ def «unsupported nesting, predicate» (xs : PList Nat) : True := | y:::ys => match ys with | PList.nil => True.intro - | _:::_:::zs => «unsupported nesting, predicate» zs - | zs => «unsupported nesting, predicate» ys + | _:::_:::zs => «unsupported nesting, predicate» zs + | zs => «unsupported nesting, predicate» ys def f1 (xs : List Nat) : Nat := @@ -221,4 +221,4 @@ match xs with | x:::xs => match xs with | PList.nil => True.intro - | _ => pf1 xs + | _ => pf1 xs diff --git a/tests/lean/run/tactic1.lean b/tests/lean/run/tactic1.lean index 8f3ed2fbed15..21b08aab4765 100644 --- a/tests/lean/run/tactic1.lean +++ b/tests/lean/run/tactic1.lean @@ -1,10 +1,10 @@ -theorem ex1 (x : Nat) (y : { v // v > x }) (z : Nat) : Nat := +def ex1 (x : Nat) (y : { v // v > x }) (z : Nat) : Nat := by { clear y x; exact z } -theorem ex2 (x : Nat) (y : { v // v > x }) (z : Nat) : Nat := +def ex2 (x : Nat) (y : { v // v > x }) (z : Nat) : Nat := by { clear x y; exact z diff --git a/tests/lean/run/thmIsProp.lean b/tests/lean/run/thmIsProp.lean new file mode 100644 index 000000000000..4d5654a63d01 --- /dev/null +++ b/tests/lean/run/thmIsProp.lean @@ -0,0 +1,57 @@ +import Lean + +open Lean + +/-- +error: (kernel) type of theorem 'bad' is not a proposition + Nat +-/ +#guard_msgs (error) in +run_meta do + addDecl <| .thmDecl { + name := `bad + levelParams := [] + type := mkConst ``Nat + value := toExpr 10 + } + +theorem foo : 10 = 10 := rfl +where aux : Nat := 20 + +/-- +info: def foo.aux : Nat := +20 +-/ +#guard_msgs in +#print foo.aux + + +theorem foo2 : 10 = 10 := + let rec aux (x : Nat) : Nat := x + 1 + rfl + +/-- +info: def foo2.aux : Nat → Nat := +fun x => x + 1 +-/ +#guard_msgs in +#print foo2.aux + + +/-- +error: type of theorem 'ugly' is not a proposition + Nat +-/ +#guard_msgs (error) in +theorem ugly : Nat := 10 + +/-- +error: type of theorem 'g' is not a proposition + Nat → Nat +-/ +#guard_msgs (error) in +mutual +theorem f (x : Nat) : x = x := rfl + +theorem g (x : Nat) : Nat := x + 1 +end From 0f193326189a35efd1b574d45ede9fd649b74f7d Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Wed, 13 Mar 2024 11:45:44 -0700 Subject: [PATCH 25/32] chore: update stage0 --- stage0/src/kernel/environment.cpp | 2 + stage0/src/kernel/kernel_exception.h | 33 +- stage0/stdlib/Init/Data/Int.c | 6 +- stage0/stdlib/Init/Data/Int/DivModLemmas.c | 6 +- stage0/stdlib/Init/Data/Int/Gcd.c | 35 +- stage0/stdlib/Init/Data/Int/Order.c | 6 +- stage0/stdlib/Init/Data/Int/Pow.c | 29 + stage0/stdlib/Init/Data/Nat/Bitwise/Lemmas.c | 6 +- stage0/stdlib/Init/Data/Nat/Div.c | 11 + stage0/stdlib/Init/Data/Nat/Dvd.c | 17 +- stage0/stdlib/Init/Data/Nat/Lemmas.c | 6 +- stage0/stdlib/Init/Data/Option/Basic.c | 60 +- stage0/stdlib/Init/Data/Option/Instances.c | 4 +- stage0/stdlib/Init/Data/Random.c | 6 +- stage0/stdlib/Init/Data/Sum.c | 69 +- stage0/stdlib/Init/MetaTypes.c | 114 +- stage0/stdlib/Init/Omega/Constraint.c | 119 +- stage0/stdlib/Init/Omega/Int.c | 6 +- stage0/stdlib/Init/Omega/IntList.c | 6 +- stage0/stdlib/Init/Omega/LinearCombo.c | 89 +- stage0/stdlib/Lean/Compiler/CSimpAttr.c | 6 +- .../stdlib/Lean/Compiler/LCNF/AuxDeclCache.c | 14 +- stage0/stdlib/Lean/Compiler/LCNF/Basic.c | 38 +- .../stdlib/Lean/Compiler/LCNF/Simp/JpCases.c | 202 +- stage0/stdlib/Lean/Compiler/LCNF/Testing.c | 4 +- stage0/stdlib/Lean/CoreM.c | 4 +- stage0/stdlib/Lean/Data/Json/Parser.c | 427 ++- stage0/stdlib/Lean/Data/Position.c | 4 +- stage0/stdlib/Lean/Elab/Command.c | 4 +- stage0/stdlib/Lean/Elab/GuardMsgs.c | 3125 +++++++++-------- stage0/stdlib/Lean/Elab/MutualDef.c | 789 +++-- stage0/stdlib/Lean/Elab/Tactic/Omega/Core.c | 176 +- .../stdlib/Lean/Elab/Tactic/Omega/MinNatAbs.c | 16 +- stage0/stdlib/Lean/Elab/Tactic/Omega/OmegaM.c | 75 +- stage0/stdlib/Lean/Elab/Tactic/SimpTrace.c | 615 ++-- stage0/stdlib/Lean/Environment.c | 506 +-- stage0/stdlib/Lean/Linter/MissingDocs.c | 4 +- stage0/stdlib/Lean/Message.c | 163 +- stage0/stdlib/Lean/Meta/KAbstract.c | 18 +- stage0/stdlib/Lean/Meta/Match/Match.c | 42 +- stage0/stdlib/Lean/Meta/Tactic/ElimInfo.c | 42 +- .../Lean/Meta/Tactic/LinearArith/Solver.c | 117 +- stage0/stdlib/Lean/Meta/Tactic/Rewrite.c | 8 + stage0/stdlib/Lean/Parser/Command.c | 4 +- stage0/stdlib/Lean/Server/References.c | 42 +- stage0/stdlib/Lean/Server/Requests.c | 4 +- stage0/stdlib/Lean/Server/Watchdog.c | 4 +- 47 files changed, 3728 insertions(+), 3355 deletions(-) create mode 100644 stage0/stdlib/Init/Data/Int/Pow.c diff --git a/stage0/src/kernel/environment.cpp b/stage0/src/kernel/environment.cpp index 56b67af4a7c2..792f4880dfe3 100644 --- a/stage0/src/kernel/environment.cpp +++ b/stage0/src/kernel/environment.cpp @@ -178,6 +178,8 @@ environment environment::add_theorem(declaration const & d, bool check) const { if (check) { // TODO(Leo): we must add support for handling tasks here type_checker checker(*this); + if (!checker.is_prop(v.get_type())) + throw theorem_type_is_not_prop(*this, v.get_name(), v.get_type()); check_constant_val(*this, v.to_constant_val(), checker); check_no_metavar_no_fvar(*this, v.get_name(), v.get_value()); expr val_type = checker.check(v.get_value(), v.get_lparams()); diff --git a/stage0/src/kernel/kernel_exception.h b/stage0/src/kernel/kernel_exception.h index 930185828fbd..793cf6403277 100644 --- a/stage0/src/kernel/kernel_exception.h +++ b/stage0/src/kernel/kernel_exception.h @@ -66,6 +66,16 @@ class declaration_has_free_vars_exception : public kernel_exception { expr const & get_expr() const { return m_expr; } }; +class theorem_type_is_not_prop : public kernel_exception { + name m_name; + expr m_type; +public: + theorem_type_is_not_prop(environment const & env, name const & n, expr const & type): + kernel_exception(env), m_name(n), m_type(type) {} + name const & get_decl_name() const { return m_name; } + expr const & get_type() const { return m_type; } +}; + class kernel_exception_with_lctx : public kernel_exception { local_ctx m_lctx; public: @@ -185,21 +195,24 @@ object * catch_kernel_exceptions(std::function const & f) { } catch (invalid_proj_exception & ex) { // 10 | invalidProj (env : Environment) (lctx : LocalContext) (proj : Expr) return mk_cnstr(0, mk_cnstr(10, ex.env(), ex.get_local_ctx(), ex.get_proj())).steal(); + } catch (theorem_type_is_not_prop & ex) { + // 11 | thmTypeIsNotProp (env : Environment) (name : Name) (type : Expr) + return mk_cnstr(0, mk_cnstr(11, ex.env(), ex.get_decl_name(), ex.get_type())).steal(); } catch (exception & ex) { - // 11 | other (msg : String) - return mk_cnstr(0, mk_cnstr(11, string_ref(ex.what()))).steal(); + // 12 | other (msg : String) + return mk_cnstr(0, mk_cnstr(12, string_ref(ex.what()))).steal(); } catch (heartbeat_exception & ex) { - // 12 | deterministicTimeout - return mk_cnstr(0, box(12)).steal(); - } catch (memory_exception & ex) { - // 13 | excessiveMemory + // 13 | deterministicTimeout return mk_cnstr(0, box(13)).steal(); - } catch (stack_space_exception & ex) { - // 14 | deepRecursion + } catch (memory_exception & ex) { + // 14 | excessiveMemory return mk_cnstr(0, box(14)).steal(); - } catch (interrupted & ex) { - // 15 | interrupted + } catch (stack_space_exception & ex) { + // 15 | deepRecursion return mk_cnstr(0, box(15)).steal(); + } catch (interrupted & ex) { + // 16 | interrupted + return mk_cnstr(0, box(16)).steal(); } } } diff --git a/stage0/stdlib/Init/Data/Int.c b/stage0/stdlib/Init/Data/Int.c index 0c748e32b669..4b39428c5af7 100644 --- a/stage0/stdlib/Init/Data/Int.c +++ b/stage0/stdlib/Init/Data/Int.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Init.Data.Int -// Imports: Init.Data.Int.Basic Init.Data.Int.Bitwise Init.Data.Int.DivMod Init.Data.Int.DivModLemmas Init.Data.Int.Gcd Init.Data.Int.Lemmas Init.Data.Int.Order +// Imports: Init.Data.Int.Basic Init.Data.Int.Bitwise Init.Data.Int.DivMod Init.Data.Int.DivModLemmas Init.Data.Int.Gcd Init.Data.Int.Lemmas Init.Data.Int.Order Init.Data.Int.Pow #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -20,6 +20,7 @@ lean_object* initialize_Init_Data_Int_DivModLemmas(uint8_t builtin, lean_object* lean_object* initialize_Init_Data_Int_Gcd(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_Int_Lemmas(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_Int_Order(uint8_t builtin, lean_object*); +lean_object* initialize_Init_Data_Int_Pow(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Init_Data_Int(uint8_t builtin, lean_object* w) { lean_object * res; @@ -46,6 +47,9 @@ lean_dec_ref(res); res = initialize_Init_Data_Int_Order(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Init_Data_Int_Pow(builtin, 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)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Init/Data/Int/DivModLemmas.c b/stage0/stdlib/Init/Data/Int/DivModLemmas.c index 05479ca826a3..5acc777a018e 100644 --- a/stage0/stdlib/Init/Data/Int/DivModLemmas.c +++ b/stage0/stdlib/Init/Data/Int/DivModLemmas.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Init.Data.Int.DivModLemmas -// Imports: Init.Data.Int.DivMod Init.Data.Int.Order Init.Data.Nat.Dvd Init.RCases +// Imports: Init.Data.Int.DivMod Init.Data.Int.Order Init.Data.Nat.Dvd #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -53,7 +53,6 @@ return x_4; lean_object* initialize_Init_Data_Int_DivMod(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_Int_Order(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_Nat_Dvd(uint8_t builtin, lean_object*); -lean_object* initialize_Init_RCases(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Init_Data_Int_DivModLemmas(uint8_t builtin, lean_object* w) { lean_object * res; @@ -68,9 +67,6 @@ lean_dec_ref(res); res = initialize_Init_Data_Nat_Dvd(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -res = initialize_Init_RCases(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); l_Int_decidableDvd___closed__1 = _init_l_Int_decidableDvd___closed__1(); lean_mark_persistent(l_Int_decidableDvd___closed__1); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Init/Data/Int/Gcd.c b/stage0/stdlib/Init/Data/Int/Gcd.c index 5597bcc29197..ca0e08436053 100644 --- a/stage0/stdlib/Init/Data/Int/Gcd.c +++ b/stage0/stdlib/Init/Data/Int/Gcd.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Init.Data.Int.Gcd -// Imports: Init.Data.Int.Basic Init.Data.Nat.Gcd +// Imports: Init.Data.Int.Basic Init.Data.Nat.Gcd Init.Data.Nat.Lcm Init.Data.Int.DivModLemmas #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -14,8 +14,11 @@ extern "C" { #endif lean_object* lean_nat_gcd(lean_object*, lean_object*); +lean_object* l_Nat_lcm(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Int_gcd(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Int_lcm(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Int_gcd___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Int_lcm___boxed(lean_object*, lean_object*); lean_object* lean_nat_abs(lean_object*); LEAN_EXPORT lean_object* l_Int_gcd(lean_object* x_1, lean_object* x_2) { _start: @@ -39,8 +42,32 @@ lean_dec(x_1); return x_3; } } +LEAN_EXPORT lean_object* l_Int_lcm(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = lean_nat_abs(x_1); +x_4 = lean_nat_abs(x_2); +x_5 = l_Nat_lcm(x_3, x_4); +lean_dec(x_4); +lean_dec(x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Int_lcm___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Int_lcm(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} lean_object* initialize_Init_Data_Int_Basic(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_Nat_Gcd(uint8_t builtin, lean_object*); +lean_object* initialize_Init_Data_Nat_Lcm(uint8_t builtin, lean_object*); +lean_object* initialize_Init_Data_Int_DivModLemmas(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Init_Data_Int_Gcd(uint8_t builtin, lean_object* w) { lean_object * res; @@ -52,6 +79,12 @@ lean_dec_ref(res); res = initialize_Init_Data_Nat_Gcd(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Init_Data_Nat_Lcm(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Init_Data_Int_DivModLemmas(builtin, 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)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Init/Data/Int/Order.c b/stage0/stdlib/Init/Data/Int/Order.c index b25c38265e70..647f8c18f00e 100644 --- a/stage0/stdlib/Init/Data/Int/Order.c +++ b/stage0/stdlib/Init/Data/Int/Order.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Init.Data.Int.Order -// Imports: Init.Data.Int.Lemmas Init.ByCases +// Imports: Init.Data.Int.Lemmas Init.ByCases Init.RCases #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -43,6 +43,7 @@ return lean_box(0); } lean_object* initialize_Init_Data_Int_Lemmas(uint8_t builtin, lean_object*); lean_object* initialize_Init_ByCases(uint8_t builtin, lean_object*); +lean_object* initialize_Init_RCases(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Init_Data_Int_Order(uint8_t builtin, lean_object* w) { lean_object * res; @@ -54,6 +55,9 @@ lean_dec_ref(res); res = initialize_Init_ByCases(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Init_RCases(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); l_Int_instTransIntLeInstLEInt = _init_l_Int_instTransIntLeInstLEInt(); l_Int_instTransIntLtInstLTIntLeInstLEInt = _init_l_Int_instTransIntLtInstLTIntLeInstLEInt(); l_Int_instTransIntLeInstLEIntLtInstLTInt = _init_l_Int_instTransIntLeInstLEIntLtInstLTInt(); diff --git a/stage0/stdlib/Init/Data/Int/Pow.c b/stage0/stdlib/Init/Data/Int/Pow.c new file mode 100644 index 000000000000..99ef7152f80c --- /dev/null +++ b/stage0/stdlib/Init/Data/Int/Pow.c @@ -0,0 +1,29 @@ +// Lean compiler output +// Module: Init.Data.Int.Pow +// Imports: Init.Data.Int.Lemmas +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +lean_object* initialize_Init_Data_Int_Lemmas(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Init_Data_Int_Pow(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Init_Data_Int_Lemmas(builtin, 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)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Init/Data/Nat/Bitwise/Lemmas.c b/stage0/stdlib/Init/Data/Nat/Bitwise/Lemmas.c index b360844dd0b8..d5efde9aa861 100644 --- a/stage0/stdlib/Init/Data/Nat/Bitwise/Lemmas.c +++ b/stage0/stdlib/Init/Data/Nat/Bitwise/Lemmas.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Init.Data.Nat.Bitwise.Lemmas -// Imports: Init.Data.Bool Init.Data.Nat.Bitwise.Basic Init.Data.Nat.Lemmas Init.TacticsExtra Init.Omega +// Imports: Init.Data.Bool Init.Data.Int.Pow Init.Data.Nat.Bitwise.Basic Init.Data.Nat.Lemmas Init.TacticsExtra Init.Omega #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -14,6 +14,7 @@ extern "C" { #endif lean_object* initialize_Init_Data_Bool(uint8_t builtin, lean_object*); +lean_object* initialize_Init_Data_Int_Pow(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_Nat_Bitwise_Basic(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_Nat_Lemmas(uint8_t builtin, lean_object*); lean_object* initialize_Init_TacticsExtra(uint8_t builtin, lean_object*); @@ -26,6 +27,9 @@ _G_initialized = true; res = initialize_Init_Data_Bool(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Init_Data_Int_Pow(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); res = initialize_Init_Data_Nat_Bitwise_Basic(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); diff --git a/stage0/stdlib/Init/Data/Nat/Div.c b/stage0/stdlib/Init/Data/Nat/Div.c index 92ae6a59bbd3..83578432be96 100644 --- a/stage0/stdlib/Init/Data/Nat/Div.c +++ b/stage0/stdlib/Init/Data/Nat/Div.c @@ -18,11 +18,20 @@ LEAN_EXPORT lean_object* l_Nat_modCore___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_instModNat; LEAN_EXPORT lean_object* l_Nat_div___boxed(lean_object*, lean_object*); lean_object* lean_nat_div(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Nat_instDvdNat; lean_object* lean_nat_mod(lean_object*, lean_object*); lean_object* lean_nat_mod(lean_object*, lean_object*); static lean_object* l_Nat_instModNat___closed__1; static lean_object* l_Nat_instDivNat___closed__1; LEAN_EXPORT lean_object* l_Nat_mod___boxed(lean_object*, lean_object*); +static lean_object* _init_l_Nat_instDvdNat() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} LEAN_EXPORT lean_object* l_Nat_div___boxed(lean_object* x_1, lean_object* x_2) { _start: { @@ -102,6 +111,8 @@ lean_dec_ref(res); res = initialize_Init_Data_Nat_Basic(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +l_Nat_instDvdNat = _init_l_Nat_instDvdNat(); +lean_mark_persistent(l_Nat_instDvdNat); l_Nat_instDivNat___closed__1 = _init_l_Nat_instDivNat___closed__1(); lean_mark_persistent(l_Nat_instDivNat___closed__1); l_Nat_instDivNat = _init_l_Nat_instDivNat(); diff --git a/stage0/stdlib/Init/Data/Nat/Dvd.c b/stage0/stdlib/Init/Data/Nat/Dvd.c index 405bfab1e2d3..034558450e8c 100644 --- a/stage0/stdlib/Init/Data/Nat/Dvd.c +++ b/stage0/stdlib/Init/Data/Nat/Dvd.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Init.Data.Nat.Dvd -// Imports: Init.Data.Nat.Div Init.TacticsExtra +// Imports: Init.Data.Nat.Div Init.Meta #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -15,17 +15,8 @@ extern "C" { #endif LEAN_EXPORT lean_object* l_Nat_decidable__dvd___boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Nat_decidable__dvd(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Nat_instDvdNat; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); lean_object* lean_nat_mod(lean_object*, lean_object*); -static lean_object* _init_l_Nat_instDvdNat() { -_start: -{ -lean_object* x_1; -x_1 = lean_box(0); -return x_1; -} -} LEAN_EXPORT uint8_t l_Nat_decidable__dvd(lean_object* x_1, lean_object* x_2) { _start: { @@ -49,7 +40,7 @@ return x_4; } } lean_object* initialize_Init_Data_Nat_Div(uint8_t builtin, lean_object*); -lean_object* initialize_Init_TacticsExtra(uint8_t builtin, lean_object*); +lean_object* initialize_Init_Meta(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Init_Data_Nat_Dvd(uint8_t builtin, lean_object* w) { lean_object * res; @@ -58,11 +49,9 @@ _G_initialized = true; res = initialize_Init_Data_Nat_Div(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -res = initialize_Init_TacticsExtra(builtin, lean_io_mk_world()); +res = initialize_Init_Meta(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l_Nat_instDvdNat = _init_l_Nat_instDvdNat(); -lean_mark_persistent(l_Nat_instDvdNat); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Init/Data/Nat/Lemmas.c b/stage0/stdlib/Init/Data/Nat/Lemmas.c index 5c16deb7f796..b9491928ce9a 100644 --- a/stage0/stdlib/Init/Data/Nat/Lemmas.c +++ b/stage0/stdlib/Init/Data/Nat/Lemmas.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Init.Data.Nat.Lemmas -// Imports: Init.Data.Nat.Dvd Init.Data.Nat.MinMax Init.Data.Nat.Log2 Init.Data.Nat.Power2 Init.Omega +// Imports: Init.Data.Nat.MinMax Init.Data.Nat.Log2 Init.Data.Nat.Power2 Init.Omega #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -269,7 +269,6 @@ lean_dec(x_2); return x_3; } } -lean_object* initialize_Init_Data_Nat_Dvd(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_Nat_MinMax(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_Nat_Log2(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_Nat_Power2(uint8_t builtin, lean_object*); @@ -279,9 +278,6 @@ LEAN_EXPORT lean_object* initialize_Init_Data_Nat_Lemmas(uint8_t builtin, lean_o lean_object * res; if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); _G_initialized = true; -res = initialize_Init_Data_Nat_Dvd(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); res = initialize_Init_Data_Nat_MinMax(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); diff --git a/stage0/stdlib/Init/Data/Option/Basic.c b/stage0/stdlib/Init/Data/Option/Basic.c index 74ad875ba9c7..e6d03c88e1a7 100644 --- a/stage0/stdlib/Init/Data/Option/Basic.c +++ b/stage0/stdlib/Init/Data/Option/Basic.c @@ -117,7 +117,6 @@ LEAN_EXPORT lean_object* l_instMonadOption___lambda__5(lean_object*, lean_object static lean_object* l_instMonadExceptOfUnitOption___closed__2; LEAN_EXPORT lean_object* l_Option_mapA___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Option_isSome(lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____at_Option_instDecidableEqOption___spec__1(lean_object*); static lean_object* l_instFunctorOption___closed__2; LEAN_EXPORT lean_object* l_Option_min(lean_object*); LEAN_EXPORT lean_object* l_instMonadExceptOfUnitOption; @@ -134,7 +133,6 @@ LEAN_EXPORT lean_object* l_Option_isSome___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Option_liftOrGet(lean_object*); static lean_object* l_Option_toArray___rarg___closed__1; LEAN_EXPORT lean_object* l_Option_merge___rarg(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____at_Option_instDecidableEqOption___spec__1___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_liftOption___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_instMonadExceptOfUnitOption___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Option_toMonad___rarg(lean_object*, lean_object*); @@ -197,67 +195,11 @@ x_2 = lean_alloc_closure((void*)(l___private_Init_Data_Option_Basic_0__Option_de return x_2; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____at_Option_instDecidableEqOption___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -if (lean_obj_tag(x_2) == 0) -{ -lean_dec(x_1); -if (lean_obj_tag(x_3) == 0) -{ -uint8_t x_4; lean_object* x_5; -x_4 = 1; -x_5 = lean_box(x_4); -return x_5; -} -else -{ -uint8_t x_6; lean_object* x_7; -lean_dec(x_3); -x_6 = 0; -x_7 = lean_box(x_6); -return x_7; -} -} -else -{ -if (lean_obj_tag(x_3) == 0) -{ -uint8_t x_8; lean_object* x_9; -lean_dec(x_2); -lean_dec(x_1); -x_8 = 0; -x_9 = lean_box(x_8); -return x_9; -} -else -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = lean_ctor_get(x_2, 0); -lean_inc(x_10); -lean_dec(x_2); -x_11 = lean_ctor_get(x_3, 0); -lean_inc(x_11); -lean_dec(x_3); -x_12 = lean_apply_2(x_1, x_10, x_11); -return x_12; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____at_Option_instDecidableEqOption___spec__1(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____at_Option_instDecidableEqOption___spec__1___rarg), 3, 0); -return x_2; -} -} LEAN_EXPORT lean_object* l_Option_instDecidableEqOption___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____at_Option_instDecidableEqOption___spec__1___rarg(x_1, x_2, x_3); +x_4 = l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____rarg(x_1, x_2, x_3); return x_4; } } diff --git a/stage0/stdlib/Init/Data/Option/Instances.c b/stage0/stdlib/Init/Data/Option/Instances.c index e5dcbc608127..ac8ca25712ae 100644 --- a/stage0/stdlib/Init/Data/Option/Instances.c +++ b/stage0/stdlib/Init/Data/Option/Instances.c @@ -28,11 +28,11 @@ LEAN_EXPORT lean_object* l_Option_instForIn_x27OptionInferInstanceMembershipInst LEAN_EXPORT lean_object* l_Option_pbind(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Option_forM___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Option_decidable__eq__none___rarg___boxed(lean_object*); +lean_object* l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Option_instForIn_x27OptionInferInstanceMembershipInstMembershipOption(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Option_instForIn_x27OptionInferInstanceMembershipInstMembershipOption___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Option_pmap(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Option_instForMOption___rarg(lean_object*, lean_object*, lean_object*); -lean_object* l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____at_Option_instDecidableEqOption___spec__1___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Option_pmap___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Option_pbind___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Option_instMembershipOption(lean_object* x_1) { @@ -49,7 +49,7 @@ LEAN_EXPORT lean_object* l_Option_instDecidableMemOptionInstMembershipOption___r lean_object* x_4; lean_object* x_5; x_4 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_4, 0, x_2); -x_5 = l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____at_Option_instDecidableEqOption___spec__1___rarg(x_1, x_3, x_4); +x_5 = l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____rarg(x_1, x_3, x_4); return x_5; } } diff --git a/stage0/stdlib/Init/Data/Random.c b/stage0/stdlib/Init/Data/Random.c index fb725bc7ade6..f61cf4ad2ff6 100644 --- a/stage0/stdlib/Init/Data/Random.c +++ b/stage0/stdlib/Init/Data/Random.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Init.Data.Random -// Imports: Init.System.IO Init.Data.Int +// Imports: Init.System.IO #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -1801,7 +1801,6 @@ return x_4; } } lean_object* initialize_Init_System_IO(uint8_t builtin, lean_object*); -lean_object* initialize_Init_Data_Int(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Init_Data_Random(uint8_t builtin, lean_object* w) { lean_object * res; @@ -1810,9 +1809,6 @@ _G_initialized = true; res = initialize_Init_System_IO(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -res = initialize_Init_Data_Int(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); l_instInhabitedStdGen___closed__1 = _init_l_instInhabitedStdGen___closed__1(); lean_mark_persistent(l_instInhabitedStdGen___closed__1); l_instInhabitedStdGen = _init_l_instInhabitedStdGen(); diff --git a/stage0/stdlib/Init/Data/Sum.c b/stage0/stdlib/Init/Data/Sum.c index ed3e337d6219..2418399b28b1 100644 --- a/stage0/stdlib/Init/Data/Sum.c +++ b/stage0/stdlib/Init/Data/Sum.c @@ -17,14 +17,12 @@ LEAN_EXPORT lean_object* l_Sum_getRight_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Sum_0__Sum_beqSum____x40_Init_Data_Sum___hyg_241____rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Sum_getRight_x3f___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Sum_0__Sum_decEqSum____x40_Init_Data_Sum___hyg_4____rarg(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Sum_0__Sum_decEqSum____x40_Init_Data_Sum___hyg_4____at_Sum_instDecidableEqSum___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Sum_getLeft_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Sum_getRight_x3f___rarg(lean_object*); LEAN_EXPORT lean_object* l_Sum_getLeft_x3f___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Sum_instDecidableEqSum(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Sum_0__Sum_decEqSum____x40_Init_Data_Sum___hyg_4_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Sum_instBEqSum___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Sum_0__Sum_decEqSum____x40_Init_Data_Sum___hyg_4____at_Sum_instDecidableEqSum___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Sum_getLeft_x3f___rarg(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Sum_0__Sum_beqSum____x40_Init_Data_Sum___hyg_241_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Sum_instDecidableEqSum___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -94,76 +92,11 @@ x_3 = lean_alloc_closure((void*)(l___private_Init_Data_Sum_0__Sum_decEqSum____x4 return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Sum_0__Sum_decEqSum____x40_Init_Data_Sum___hyg_4____at_Sum_instDecidableEqSum___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -if (lean_obj_tag(x_3) == 0) -{ -lean_dec(x_2); -if (lean_obj_tag(x_4) == 0) -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_5 = lean_ctor_get(x_3, 0); -lean_inc(x_5); -lean_dec(x_3); -x_6 = lean_ctor_get(x_4, 0); -lean_inc(x_6); -lean_dec(x_4); -x_7 = lean_apply_2(x_1, x_5, x_6); -return x_7; -} -else -{ -uint8_t x_8; lean_object* x_9; -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -x_8 = 0; -x_9 = lean_box(x_8); -return x_9; -} -} -else -{ -lean_dec(x_1); -if (lean_obj_tag(x_4) == 0) -{ -uint8_t x_10; lean_object* x_11; -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_10 = 0; -x_11 = lean_box(x_10); -return x_11; -} -else -{ -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 = lean_ctor_get(x_4, 0); -lean_inc(x_13); -lean_dec(x_4); -x_14 = lean_apply_2(x_2, x_12, x_13); -return x_14; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Init_Data_Sum_0__Sum_decEqSum____x40_Init_Data_Sum___hyg_4____at_Sum_instDecidableEqSum___spec__1(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l___private_Init_Data_Sum_0__Sum_decEqSum____x40_Init_Data_Sum___hyg_4____at_Sum_instDecidableEqSum___spec__1___rarg), 4, 0); -return x_3; -} -} LEAN_EXPORT lean_object* l_Sum_instDecidableEqSum___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___private_Init_Data_Sum_0__Sum_decEqSum____x40_Init_Data_Sum___hyg_4____at_Sum_instDecidableEqSum___spec__1___rarg(x_1, x_2, x_3, x_4); +x_5 = l___private_Init_Data_Sum_0__Sum_decEqSum____x40_Init_Data_Sum___hyg_4____rarg(x_1, x_2, x_3, x_4); return x_5; } } diff --git a/stage0/stdlib/Init/MetaTypes.c b/stage0/stdlib/Init/MetaTypes.c index 3163a1ffda1e..0e1b58483672 100644 --- a/stage0/stdlib/Init/MetaTypes.c +++ b/stage0/stdlib/Init/MetaTypes.c @@ -24,7 +24,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_TransparencyMode_toCtorIdx___boxed(lean_obj LEAN_EXPORT uint8_t l_Lean_Meta_DSimp_Config_zeta___default; LEAN_EXPORT lean_object* l_Lean_NameGenerator_idx___default; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_Config_maxDischargeDepth___default; -uint8_t l_List_hasDecEq___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_MetaTypes_0__Lean_Meta_Simp_beqConfig____x40_Init_MetaTypes___hyg_660____boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Init_MetaTypes_0__Lean_Meta_Simp_beqConfig____x40_Init_MetaTypes___hyg_660_(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_Simp_Config_autoUnfold___default; @@ -42,6 +41,7 @@ LEAN_EXPORT uint8_t l_Lean_Meta_Simp_Config_contextual___default; static lean_object* l_Lean_Meta_Simp_instInhabitedConfig___closed__1; LEAN_EXPORT uint8_t l_Lean_Meta_DSimp_Config_etaStruct___default; LEAN_EXPORT uint8_t l_Lean_Meta_Simp_Config_memoize___default; +LEAN_EXPORT uint8_t l_List_hasDecEq___at___private_Init_MetaTypes_0__Lean_Meta_beqOccurrences____x40_Init_MetaTypes___hyg_1066____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_NameGenerator_namePrefix___default___closed__1; LEAN_EXPORT uint8_t l___private_Init_MetaTypes_0__Lean_Meta_beqOccurrences____x40_Init_MetaTypes___hyg_1066_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_TransparencyMode_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); @@ -67,6 +67,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_instBEqOccurrences; LEAN_EXPORT uint8_t l_Lean_Meta_Simp_Config_decide___default; LEAN_EXPORT lean_object* l___private_Init_MetaTypes_0__Lean_Meta_beqTransparencyMode____x40_Init_MetaTypes___hyg_67____boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_Simp_Config_dsimp___default; +uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_Simp_Config_zetaDelta___default; LEAN_EXPORT uint8_t l_Lean_Meta_Simp_Config_etaStruct___default; LEAN_EXPORT lean_object* l_Lean_Meta_instBEqEtaStructMode; @@ -77,11 +78,10 @@ LEAN_EXPORT uint8_t l_Lean_Meta_Simp_Config_singlePass___default; LEAN_EXPORT uint8_t l___private_Init_MetaTypes_0__Lean_Meta_DSimp_beqConfig____x40_Init_MetaTypes___hyg_233_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_instInhabitedOccurrences; LEAN_EXPORT lean_object* l_Lean_NameGenerator_namePrefix___default; +LEAN_EXPORT lean_object* l_List_hasDecEq___at___private_Init_MetaTypes_0__Lean_Meta_beqOccurrences____x40_Init_MetaTypes___hyg_1066____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Init_MetaTypes_0__Lean_Meta_beqTransparencyMode____x40_Init_MetaTypes___hyg_67_(uint8_t, uint8_t); -static lean_object* l___private_Init_MetaTypes_0__Lean_Meta_beqOccurrences____x40_Init_MetaTypes___hyg_1066____closed__1; LEAN_EXPORT uint8_t l_Lean_Meta_Simp_Config_ground___default; LEAN_EXPORT uint8_t l_Lean_Meta_DSimp_Config_decide___default; -lean_object* l_Nat_decEq___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_instBEqTransparencyMode; LEAN_EXPORT uint8_t l_Lean_Meta_Simp_Config_eta___default; LEAN_EXPORT uint8_t l_Lean_Meta_DSimp_Config_beta___default; @@ -1858,12 +1858,54 @@ x_1 = lean_box(0); return x_1; } } -static lean_object* _init_l___private_Init_MetaTypes_0__Lean_Meta_beqOccurrences____x40_Init_MetaTypes___hyg_1066____closed__1() { +LEAN_EXPORT uint8_t l_List_hasDecEq___at___private_Init_MetaTypes_0__Lean_Meta_beqOccurrences____x40_Init_MetaTypes___hyg_1066____spec__1(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Nat_decEq___boxed), 2, 0); -return x_1; +if (lean_obj_tag(x_1) == 0) +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_3; +x_3 = 1; +return x_3; +} +else +{ +uint8_t x_4; +x_4 = 0; +return x_4; +} +} +else +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_5; +x_5 = 0; +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_6 = lean_ctor_get(x_1, 0); +x_7 = lean_ctor_get(x_1, 1); +x_8 = lean_ctor_get(x_2, 0); +x_9 = lean_ctor_get(x_2, 1); +x_10 = lean_nat_dec_eq(x_6, x_8); +if (x_10 == 0) +{ +uint8_t x_11; +x_11 = 0; +return x_11; +} +else +{ +x_1 = x_7; +x_2 = x_9; +goto _start; +} +} +} } } LEAN_EXPORT uint8_t l___private_Init_MetaTypes_0__Lean_Meta_beqOccurrences____x40_Init_MetaTypes___hyg_1066_(lean_object* x_1, lean_object* x_2) { @@ -1881,7 +1923,6 @@ return x_3; else { uint8_t x_4; -lean_dec(x_2); x_4 = 0; return x_4; } @@ -1890,58 +1931,57 @@ case 1: { if (lean_obj_tag(x_2) == 1) { -lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; +lean_object* x_5; lean_object* x_6; uint8_t x_7; x_5 = lean_ctor_get(x_1, 0); -lean_inc(x_5); -lean_dec(x_1); x_6 = lean_ctor_get(x_2, 0); -lean_inc(x_6); -lean_dec(x_2); -x_7 = l___private_Init_MetaTypes_0__Lean_Meta_beqOccurrences____x40_Init_MetaTypes___hyg_1066____closed__1; -x_8 = l_List_hasDecEq___rarg(x_7, x_5, x_6); -return x_8; +x_7 = l_List_hasDecEq___at___private_Init_MetaTypes_0__Lean_Meta_beqOccurrences____x40_Init_MetaTypes___hyg_1066____spec__1(x_5, x_6); +return x_7; } else { -uint8_t x_9; -lean_dec(x_2); -lean_dec(x_1); -x_9 = 0; -return x_9; +uint8_t x_8; +x_8 = 0; +return x_8; } } default: { if (lean_obj_tag(x_2) == 2) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_10 = lean_ctor_get(x_1, 0); -lean_inc(x_10); -lean_dec(x_1); -x_11 = lean_ctor_get(x_2, 0); -lean_inc(x_11); -lean_dec(x_2); -x_12 = l___private_Init_MetaTypes_0__Lean_Meta_beqOccurrences____x40_Init_MetaTypes___hyg_1066____closed__1; -x_13 = l_List_hasDecEq___rarg(x_12, x_10, x_11); -return x_13; +lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_9 = lean_ctor_get(x_1, 0); +x_10 = lean_ctor_get(x_2, 0); +x_11 = l_List_hasDecEq___at___private_Init_MetaTypes_0__Lean_Meta_beqOccurrences____x40_Init_MetaTypes___hyg_1066____spec__1(x_9, x_10); +return x_11; } else { -uint8_t x_14; -lean_dec(x_2); -lean_dec(x_1); -x_14 = 0; -return x_14; +uint8_t x_12; +x_12 = 0; +return x_12; +} } } } } +LEAN_EXPORT lean_object* l_List_hasDecEq___at___private_Init_MetaTypes_0__Lean_Meta_beqOccurrences____x40_Init_MetaTypes___hyg_1066____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_List_hasDecEq___at___private_Init_MetaTypes_0__Lean_Meta_beqOccurrences____x40_Init_MetaTypes___hyg_1066____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_MetaTypes_0__Lean_Meta_beqOccurrences____x40_Init_MetaTypes___hyg_1066____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; x_3 = l___private_Init_MetaTypes_0__Lean_Meta_beqOccurrences____x40_Init_MetaTypes___hyg_1066_(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); x_4 = lean_box(x_3); return x_4; } @@ -2052,8 +2092,6 @@ l_Lean_Meta_Simp_neutralConfig = _init_l_Lean_Meta_Simp_neutralConfig(); lean_mark_persistent(l_Lean_Meta_Simp_neutralConfig); l_Lean_Meta_instInhabitedOccurrences = _init_l_Lean_Meta_instInhabitedOccurrences(); lean_mark_persistent(l_Lean_Meta_instInhabitedOccurrences); -l___private_Init_MetaTypes_0__Lean_Meta_beqOccurrences____x40_Init_MetaTypes___hyg_1066____closed__1 = _init_l___private_Init_MetaTypes_0__Lean_Meta_beqOccurrences____x40_Init_MetaTypes___hyg_1066____closed__1(); -lean_mark_persistent(l___private_Init_MetaTypes_0__Lean_Meta_beqOccurrences____x40_Init_MetaTypes___hyg_1066____closed__1); l_Lean_Meta_instBEqOccurrences___closed__1 = _init_l_Lean_Meta_instBEqOccurrences___closed__1(); lean_mark_persistent(l_Lean_Meta_instBEqOccurrences___closed__1); l_Lean_Meta_instBEqOccurrences = _init_l_Lean_Meta_instBEqOccurrences(); diff --git a/stage0/stdlib/Init/Omega/Constraint.c b/stage0/stdlib/Init/Omega/Constraint.c index 89638fbe50a7..a0f7dcc56feb 100644 --- a/stage0/stdlib/Init/Omega/Constraint.c +++ b/stage0/stdlib/Init/Omega/Constraint.c @@ -16,7 +16,7 @@ extern "C" { LEAN_EXPORT lean_object* l_Option_repr___at___private_Init_Omega_Constraint_0__Lean_Omega_reprConstraint____x40_Init_Omega_Constraint___hyg_286____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Omega_Constraint_add___boxed(lean_object*, lean_object*); static lean_object* l___private_Init_Omega_Constraint_0__Lean_Omega_reprConstraint____x40_Init_Omega_Constraint___hyg_286____closed__15; -LEAN_EXPORT lean_object* l___private_Init_Omega_Constraint_0__Lean_Omega_decEqConstraint____x40_Init_Omega_Constraint___hyg_142_(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Init_Omega_Constraint_0__Lean_Omega_decEqConstraint____x40_Init_Omega_Constraint___hyg_142_(lean_object*, lean_object*); lean_object* l_List_foldr___at_Lean_Omega_IntList_gcd___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Omega_Constraint_combo(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Omega_Constraint_isImpossible___boxed(lean_object*); @@ -52,9 +52,9 @@ static lean_object* l_Option_repr___at___private_Init_Omega_Constraint_0__Lean_O LEAN_EXPORT uint8_t l_Lean_Omega_Constraint_sat(lean_object*, lean_object*); static lean_object* l_Lean_Omega_Constraint_neg___closed__1; static lean_object* l_Lean_Omega_Constraint_instToStringConstraint___closed__4; +LEAN_EXPORT lean_object* l_Lean_Omega_instDecidableEqConstraint___boxed(lean_object*, lean_object*); lean_object* l_Lean_Omega_IntList_sdiv(lean_object*, lean_object*); static lean_object* l_Lean_Omega_instBEqConstraint___closed__1; -static lean_object* l___private_Init_Omega_Constraint_0__Lean_Omega_decEqConstraint____x40_Init_Omega_Constraint___hyg_142____closed__1; LEAN_EXPORT uint8_t l___private_Init_Omega_Constraint_0__Lean_Omega_beqConstraint____x40_Init_Omega_Constraint___hyg_65_(lean_object*, lean_object*); static lean_object* l___private_Init_Omega_Constraint_0__Lean_Omega_reprConstraint____x40_Init_Omega_Constraint___hyg_286____closed__10; lean_object* lean_nat_to_int(lean_object*); @@ -83,12 +83,14 @@ LEAN_EXPORT lean_object* l___private_Init_Omega_Constraint_0__Lean_Omega_tidy_x3 lean_object* l_Int_bmod(lean_object*, lean_object*); static lean_object* l___private_Init_Omega_Constraint_0__Lean_Omega_reprConstraint____x40_Init_Omega_Constraint___hyg_286____closed__9; static lean_object* l___private_Init_Omega_Constraint_0__Lean_Omega_reprConstraint____x40_Init_Omega_Constraint___hyg_286____closed__2; -LEAN_EXPORT lean_object* l_Lean_Omega_instDecidableEqConstraint(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_Omega_instDecidableEqConstraint(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____at___private_Init_Omega_Constraint_0__Lean_Omega_decEqConstraint____x40_Init_Omega_Constraint___hyg_142____spec__1(lean_object*, lean_object*); lean_object* lean_int_sub(lean_object*, lean_object*); static lean_object* l_Lean_Omega_instReprConstraint___closed__1; LEAN_EXPORT lean_object* l_Lean_Omega_bmod__coeffs___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Omega_Constraint_sat___boxed(lean_object*, lean_object*); lean_object* lean_nat_abs(lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Omega_Constraint_0__Lean_Omega_decEqConstraint____x40_Init_Omega_Constraint___hyg_142____boxed(lean_object*, lean_object*); static lean_object* l_Lean_Omega_Constraint_instToStringConstraint___closed__7; lean_object* lean_string_length(lean_object*); static lean_object* l___private_Init_Omega_Constraint_0__Lean_Omega_reprConstraint____x40_Init_Omega_Constraint___hyg_286____closed__1; @@ -107,6 +109,7 @@ uint8_t lean_int_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Omega_Constraint_0__Lean_Omega_tidy_x3f_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Omega_bmod__div__term(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____at___private_Init_Omega_Constraint_0__Lean_Omega_decEqConstraint____x40_Init_Omega_Constraint___hyg_142____spec__1___boxed(lean_object*, lean_object*); lean_object* l_Lean_Omega_IntList_smul(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Option_repr___at___private_Init_Omega_Constraint_0__Lean_Omega_reprConstraint____x40_Init_Omega_Constraint___hyg_286____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Omega_Constraint_map(lean_object*, lean_object*); @@ -132,11 +135,9 @@ static lean_object* l_Lean_Omega_Constraint_impossible___closed__4; LEAN_EXPORT lean_object* l___private_Init_Omega_Constraint_0__Lean_Omega_tidy_x3f_match__1_splitter(lean_object*); LEAN_EXPORT lean_object* l_Lean_Omega_normalize_x3f(lean_object*); lean_object* lean_int_neg(lean_object*); -lean_object* l_Int_decEq___boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Init_Omega_Constraint_0__Lean_Omega_beqConstraint____x40_Init_Omega_Constraint___hyg_65____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Omega_Constraint_instToStringConstraint___closed__2; lean_object* lean_nat_add(lean_object*, lean_object*); -lean_object* l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____at_Option_instDecidableEqOption___spec__1___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Omega_Constraint_0__Lean_Omega_reprConstraint____x40_Init_Omega_Constraint___hyg_286____closed__13; static lean_object* l_Lean_Omega_Constraint_impossible___closed__2; static lean_object* l_Lean_Omega_Constraint_map___closed__1; @@ -301,57 +302,107 @@ x_1 = l_Lean_Omega_instBEqConstraint___closed__1; return x_1; } } -static lean_object* _init_l___private_Init_Omega_Constraint_0__Lean_Omega_decEqConstraint____x40_Init_Omega_Constraint___hyg_142____closed__1() { +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____at___private_Init_Omega_Constraint_0__Lean_Omega_decEqConstraint____x40_Init_Omega_Constraint___hyg_142____spec__1(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Int_decEq___boxed), 2, 0); -return x_1; +if (lean_obj_tag(x_1) == 0) +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_3; +x_3 = 1; +return x_3; +} +else +{ +uint8_t x_4; +x_4 = 0; +return x_4; } } -LEAN_EXPORT lean_object* l___private_Init_Omega_Constraint_0__Lean_Omega_decEqConstraint____x40_Init_Omega_Constraint___hyg_142_(lean_object* x_1, lean_object* x_2) { +else +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_5; +x_5 = 0; +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_6 = lean_ctor_get(x_1, 0); +x_7 = lean_ctor_get(x_2, 0); +x_8 = lean_int_dec_eq(x_6, x_7); +return x_8; +} +} +} +} +LEAN_EXPORT uint8_t l___private_Init_Omega_Constraint_0__Lean_Omega_decEqConstraint____x40_Init_Omega_Constraint___hyg_142_(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; +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; x_3 = lean_ctor_get(x_1, 0); -lean_inc(x_3); x_4 = lean_ctor_get(x_1, 1); -lean_inc(x_4); -lean_dec(x_1); x_5 = lean_ctor_get(x_2, 0); -lean_inc(x_5); x_6 = lean_ctor_get(x_2, 1); -lean_inc(x_6); -lean_dec(x_2); -x_7 = l___private_Init_Omega_Constraint_0__Lean_Omega_decEqConstraint____x40_Init_Omega_Constraint___hyg_142____closed__1; -x_8 = l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____at_Option_instDecidableEqOption___spec__1___rarg(x_7, x_3, x_5); -x_9 = lean_unbox(x_8); -lean_dec(x_8); -if (x_9 == 0) +x_7 = l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____at___private_Init_Omega_Constraint_0__Lean_Omega_decEqConstraint____x40_Init_Omega_Constraint___hyg_142____spec__1(x_3, x_5); +if (x_7 == 0) { -uint8_t x_10; lean_object* x_11; -lean_dec(x_6); -lean_dec(x_4); -x_10 = 0; -x_11 = lean_box(x_10); -return x_11; +uint8_t x_8; +x_8 = 0; +return x_8; } else { -lean_object* x_12; -x_12 = l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____at_Option_instDecidableEqOption___spec__1___rarg(x_7, x_4, x_6); -return x_12; +uint8_t x_9; +x_9 = l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____at___private_Init_Omega_Constraint_0__Lean_Omega_decEqConstraint____x40_Init_Omega_Constraint___hyg_142____spec__1(x_4, x_6); +return x_9; } } } -LEAN_EXPORT lean_object* l_Lean_Omega_instDecidableEqConstraint(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____at___private_Init_Omega_Constraint_0__Lean_Omega_decEqConstraint____x40_Init_Omega_Constraint___hyg_142____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_3; +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____at___private_Init_Omega_Constraint_0__Lean_Omega_decEqConstraint____x40_Init_Omega_Constraint___hyg_142____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_Omega_Constraint_0__Lean_Omega_decEqConstraint____x40_Init_Omega_Constraint___hyg_142____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Init_Omega_Constraint_0__Lean_Omega_decEqConstraint____x40_Init_Omega_Constraint___hyg_142_(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l_Lean_Omega_instDecidableEqConstraint(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; x_3 = l___private_Init_Omega_Constraint_0__Lean_Omega_decEqConstraint____x40_Init_Omega_Constraint___hyg_142_(x_1, x_2); return x_3; } } +LEAN_EXPORT lean_object* l_Lean_Omega_instDecidableEqConstraint___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_Omega_instDecidableEqConstraint(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} static lean_object* _init_l_Option_repr___at___private_Init_Omega_Constraint_0__Lean_Omega_reprConstraint____x40_Init_Omega_Constraint___hyg_286____spec__1___closed__1() { _start: { @@ -3027,8 +3078,6 @@ l_Lean_Omega_instBEqConstraint___closed__1 = _init_l_Lean_Omega_instBEqConstrain lean_mark_persistent(l_Lean_Omega_instBEqConstraint___closed__1); l_Lean_Omega_instBEqConstraint = _init_l_Lean_Omega_instBEqConstraint(); lean_mark_persistent(l_Lean_Omega_instBEqConstraint); -l___private_Init_Omega_Constraint_0__Lean_Omega_decEqConstraint____x40_Init_Omega_Constraint___hyg_142____closed__1 = _init_l___private_Init_Omega_Constraint_0__Lean_Omega_decEqConstraint____x40_Init_Omega_Constraint___hyg_142____closed__1(); -lean_mark_persistent(l___private_Init_Omega_Constraint_0__Lean_Omega_decEqConstraint____x40_Init_Omega_Constraint___hyg_142____closed__1); l_Option_repr___at___private_Init_Omega_Constraint_0__Lean_Omega_reprConstraint____x40_Init_Omega_Constraint___hyg_286____spec__1___closed__1 = _init_l_Option_repr___at___private_Init_Omega_Constraint_0__Lean_Omega_reprConstraint____x40_Init_Omega_Constraint___hyg_286____spec__1___closed__1(); lean_mark_persistent(l_Option_repr___at___private_Init_Omega_Constraint_0__Lean_Omega_reprConstraint____x40_Init_Omega_Constraint___hyg_286____spec__1___closed__1); l_Option_repr___at___private_Init_Omega_Constraint_0__Lean_Omega_reprConstraint____x40_Init_Omega_Constraint___hyg_286____spec__1___closed__2 = _init_l_Option_repr___at___private_Init_Omega_Constraint_0__Lean_Omega_reprConstraint____x40_Init_Omega_Constraint___hyg_286____spec__1___closed__2(); diff --git a/stage0/stdlib/Init/Omega/Int.c b/stage0/stdlib/Init/Omega/Int.c index c0900b6e9f24..a019bb31f5fb 100644 --- a/stage0/stdlib/Init/Omega/Int.c +++ b/stage0/stdlib/Init/Omega/Int.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Init.Omega.Int -// Imports: Init.Data.Int.DivMod Init.Data.Int.Order Init.Data.Nat.Basic +// Imports: Init.Data.Int.DivMod Init.Data.Int.Order #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -76,7 +76,6 @@ return x_4; } lean_object* initialize_Init_Data_Int_DivMod(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_Int_Order(uint8_t builtin, lean_object*); -lean_object* initialize_Init_Data_Nat_Basic(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Init_Omega_Int(uint8_t builtin, lean_object* w) { lean_object * res; @@ -88,9 +87,6 @@ lean_dec_ref(res); res = initialize_Init_Data_Int_Order(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -res = initialize_Init_Data_Nat_Basic(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); l___private_Init_Omega_Int_0__Int_neg_match__1_splitter___rarg___closed__1 = _init_l___private_Init_Omega_Int_0__Int_neg_match__1_splitter___rarg___closed__1(); lean_mark_persistent(l___private_Init_Omega_Int_0__Int_neg_match__1_splitter___rarg___closed__1); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Init/Omega/IntList.c b/stage0/stdlib/Init/Omega/IntList.c index ff26f5959359..e256f01ad478 100644 --- a/stage0/stdlib/Init/Omega/IntList.c +++ b/stage0/stdlib/Init/Omega/IntList.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Init.Omega.IntList -// Imports: Init.Data.List.Lemmas Init.Data.Int.DivModLemmas Init.Data.Int.Gcd +// Imports: Init.Data.List.Lemmas Init.Data.Int.DivModLemmas Init.Data.Nat.Gcd #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -1209,7 +1209,7 @@ return x_9; } lean_object* initialize_Init_Data_List_Lemmas(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_Int_DivModLemmas(uint8_t builtin, lean_object*); -lean_object* initialize_Init_Data_Int_Gcd(uint8_t builtin, lean_object*); +lean_object* initialize_Init_Data_Nat_Gcd(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Init_Omega_IntList(uint8_t builtin, lean_object* w) { lean_object * res; @@ -1221,7 +1221,7 @@ lean_dec_ref(res); res = initialize_Init_Data_Int_DivModLemmas(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -res = initialize_Init_Data_Int_Gcd(builtin, lean_io_mk_world()); +res = initialize_Init_Data_Nat_Gcd(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); l_Lean_Omega_IntList_get___closed__1 = _init_l_Lean_Omega_IntList_get___closed__1(); diff --git a/stage0/stdlib/Init/Omega/LinearCombo.c b/stage0/stdlib/Init/Omega/LinearCombo.c index b0516fc9daea..661cdb6c301a 100644 --- a/stage0/stdlib/Init/Omega/LinearCombo.c +++ b/stage0/stdlib/Init/Omega/LinearCombo.c @@ -15,7 +15,6 @@ extern "C" { #endif static lean_object* l_List_repr_x27___at___private_Init_Omega_LinearCombo_0__Lean_Omega_reprLinearCombo____x40_Init_Omega_LinearCombo___hyg_171____spec__1___closed__2; LEAN_EXPORT lean_object* l___private_Init_Omega_LinearCombo_0__Lean_Omega_reprLinearCombo____x40_Init_Omega_LinearCombo___hyg_171_(lean_object*, lean_object*); -static lean_object* l___private_Init_Omega_LinearCombo_0__Lean_Omega_decEqLinearCombo____x40_Init_Omega_LinearCombo___hyg_27____closed__1; static lean_object* l___private_Init_Omega_LinearCombo_0__Lean_Omega_reprLinearCombo____x40_Init_Omega_LinearCombo___hyg_171____closed__5; LEAN_EXPORT lean_object* l_Lean_Omega_LinearCombo_eval___boxed(lean_object*, lean_object*); static lean_object* l_List_mapTR_loop___at_Lean_Omega_LinearCombo_instToStringLinearCombo___spec__1___closed__1; @@ -25,7 +24,6 @@ static lean_object* l_List_repr_x27___at___private_Init_Omega_LinearCombo_0__Lea LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Omega_LinearCombo_instToStringLinearCombo___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Omega_instDecidableEqLinearCombo___boxed(lean_object*, lean_object*); static lean_object* l_List_repr_x27___at___private_Init_Omega_LinearCombo_0__Lean_Omega_reprLinearCombo____x40_Init_Omega_LinearCombo___hyg_171____spec__1___closed__1; -uint8_t l_List_hasDecEq___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Omega_LinearCombo_instNegLinearCombo; LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Init_Omega_LinearCombo_0__Lean_Omega_reprLinearCombo____x40_Init_Omega_LinearCombo___hyg_171____spec__3(lean_object*, lean_object*); lean_object* l_List_zipWithAll___rarg(lean_object*, lean_object*, lean_object*); @@ -61,6 +59,7 @@ LEAN_EXPORT lean_object* l___private_Init_Omega_LinearCombo_0__Lean_Omega_decEqL static lean_object* l___private_Init_Omega_LinearCombo_0__Lean_Omega_reprLinearCombo____x40_Init_Omega_LinearCombo___hyg_171____closed__11; lean_object* lean_int_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_repr_x27___at___private_Init_Omega_LinearCombo_0__Lean_Omega_reprLinearCombo____x40_Init_Omega_LinearCombo___hyg_171____spec__1(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_List_hasDecEq___at___private_Init_Omega_LinearCombo_0__Lean_Omega_decEqLinearCombo____x40_Init_Omega_LinearCombo___hyg_27____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Omega_LinearCombo_neg(lean_object*); static lean_object* l___private_Init_Omega_LinearCombo_0__Lean_Omega_reprLinearCombo____x40_Init_Omega_LinearCombo___hyg_171____closed__10; LEAN_EXPORT lean_object* l_Lean_Omega_LinearCombo_instSubLinearCombo; @@ -97,10 +96,10 @@ static lean_object* l_Lean_Omega_LinearCombo_add___closed__1; uint8_t lean_int_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Omega_LinearCombo_coordinate___boxed(lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_hasDecEq___at___private_Init_Omega_LinearCombo_0__Lean_Omega_decEqLinearCombo____x40_Init_Omega_LinearCombo___hyg_27____spec__1___boxed(lean_object*, lean_object*); static lean_object* l_List_repr_x27___at___private_Init_Omega_LinearCombo_0__Lean_Omega_reprLinearCombo____x40_Init_Omega_LinearCombo___hyg_171____spec__1___closed__5; LEAN_EXPORT lean_object* l_Lean_Omega_LinearCombo_mul(lean_object*, lean_object*); lean_object* lean_int_neg(lean_object*); -lean_object* l_Int_decEq___boxed(lean_object*, lean_object*); static lean_object* l___private_Init_Omega_LinearCombo_0__Lean_Omega_reprLinearCombo____x40_Init_Omega_LinearCombo___hyg_171____closed__2; lean_object* lean_nat_add(lean_object*, lean_object*); static lean_object* l_Lean_Omega_LinearCombo_instInhabitedLinearCombo___closed__1; @@ -137,12 +136,54 @@ x_1 = lean_box(0); return x_1; } } -static lean_object* _init_l___private_Init_Omega_LinearCombo_0__Lean_Omega_decEqLinearCombo____x40_Init_Omega_LinearCombo___hyg_27____closed__1() { +LEAN_EXPORT uint8_t l_List_hasDecEq___at___private_Init_Omega_LinearCombo_0__Lean_Omega_decEqLinearCombo____x40_Init_Omega_LinearCombo___hyg_27____spec__1(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Int_decEq___boxed), 2, 0); -return x_1; +if (lean_obj_tag(x_1) == 0) +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_3; +x_3 = 1; +return x_3; +} +else +{ +uint8_t x_4; +x_4 = 0; +return x_4; +} +} +else +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_5; +x_5 = 0; +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_6 = lean_ctor_get(x_1, 0); +x_7 = lean_ctor_get(x_1, 1); +x_8 = lean_ctor_get(x_2, 0); +x_9 = lean_ctor_get(x_2, 1); +x_10 = lean_int_dec_eq(x_6, x_8); +if (x_10 == 0) +{ +uint8_t x_11; +x_11 = 0; +return x_11; +} +else +{ +x_1 = x_7; +x_2 = x_9; +goto _start; +} +} +} } } LEAN_EXPORT uint8_t l___private_Init_Omega_LinearCombo_0__Lean_Omega_decEqLinearCombo____x40_Init_Omega_LinearCombo___hyg_27_(lean_object* x_1, lean_object* x_2) { @@ -150,40 +191,42 @@ LEAN_EXPORT uint8_t l___private_Init_Omega_LinearCombo_0__Lean_Omega_decEqLinear { lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; x_3 = lean_ctor_get(x_1, 0); -lean_inc(x_3); x_4 = lean_ctor_get(x_1, 1); -lean_inc(x_4); -lean_dec(x_1); x_5 = lean_ctor_get(x_2, 0); -lean_inc(x_5); x_6 = lean_ctor_get(x_2, 1); -lean_inc(x_6); -lean_dec(x_2); x_7 = lean_int_dec_eq(x_3, x_5); -lean_dec(x_5); -lean_dec(x_3); if (x_7 == 0) { uint8_t x_8; -lean_dec(x_6); -lean_dec(x_4); x_8 = 0; return x_8; } else { -lean_object* x_9; uint8_t x_10; -x_9 = l___private_Init_Omega_LinearCombo_0__Lean_Omega_decEqLinearCombo____x40_Init_Omega_LinearCombo___hyg_27____closed__1; -x_10 = l_List_hasDecEq___rarg(x_9, x_4, x_6); -return x_10; +uint8_t x_9; +x_9 = l_List_hasDecEq___at___private_Init_Omega_LinearCombo_0__Lean_Omega_decEqLinearCombo____x40_Init_Omega_LinearCombo___hyg_27____spec__1(x_4, x_6); +return x_9; } } } +LEAN_EXPORT lean_object* l_List_hasDecEq___at___private_Init_Omega_LinearCombo_0__Lean_Omega_decEqLinearCombo____x40_Init_Omega_LinearCombo___hyg_27____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_List_hasDecEq___at___private_Init_Omega_LinearCombo_0__Lean_Omega_decEqLinearCombo____x40_Init_Omega_LinearCombo___hyg_27____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_Omega_LinearCombo_0__Lean_Omega_decEqLinearCombo____x40_Init_Omega_LinearCombo___hyg_27____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; x_3 = l___private_Init_Omega_LinearCombo_0__Lean_Omega_decEqLinearCombo____x40_Init_Omega_LinearCombo___hyg_27_(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); x_4 = lean_box(x_3); return x_4; } @@ -201,6 +244,8 @@ LEAN_EXPORT lean_object* l_Lean_Omega_instDecidableEqLinearCombo___boxed(lean_ob { uint8_t x_3; lean_object* x_4; x_3 = l_Lean_Omega_instDecidableEqLinearCombo(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); x_4 = lean_box(x_3); return x_4; } @@ -1213,8 +1258,6 @@ l_Lean_Omega_LinearCombo_const___default = _init_l_Lean_Omega_LinearCombo_const_ lean_mark_persistent(l_Lean_Omega_LinearCombo_const___default); l_Lean_Omega_LinearCombo_coeffs___default = _init_l_Lean_Omega_LinearCombo_coeffs___default(); lean_mark_persistent(l_Lean_Omega_LinearCombo_coeffs___default); -l___private_Init_Omega_LinearCombo_0__Lean_Omega_decEqLinearCombo____x40_Init_Omega_LinearCombo___hyg_27____closed__1 = _init_l___private_Init_Omega_LinearCombo_0__Lean_Omega_decEqLinearCombo____x40_Init_Omega_LinearCombo___hyg_27____closed__1(); -lean_mark_persistent(l___private_Init_Omega_LinearCombo_0__Lean_Omega_decEqLinearCombo____x40_Init_Omega_LinearCombo___hyg_27____closed__1); l_List_repr_x27___at___private_Init_Omega_LinearCombo_0__Lean_Omega_reprLinearCombo____x40_Init_Omega_LinearCombo___hyg_171____spec__1___closed__1 = _init_l_List_repr_x27___at___private_Init_Omega_LinearCombo_0__Lean_Omega_reprLinearCombo____x40_Init_Omega_LinearCombo___hyg_171____spec__1___closed__1(); lean_mark_persistent(l_List_repr_x27___at___private_Init_Omega_LinearCombo_0__Lean_Omega_reprLinearCombo____x40_Init_Omega_LinearCombo___hyg_171____spec__1___closed__1); l_List_repr_x27___at___private_Init_Omega_LinearCombo_0__Lean_Omega_reprLinearCombo____x40_Init_Omega_LinearCombo___hyg_171____spec__1___closed__2 = _init_l_List_repr_x27___at___private_Init_Omega_LinearCombo_0__Lean_Omega_reprLinearCombo____x40_Init_Omega_LinearCombo___hyg_171____spec__1___closed__2(); diff --git a/stage0/stdlib/Lean/Compiler/CSimpAttr.c b/stage0/stdlib/Lean/Compiler/CSimpAttr.c index 3862c903f86c..6b8c39ffecd5 100644 --- a/stage0/stdlib/Lean/Compiler/CSimpAttr.c +++ b/stage0/stdlib/Lean/Compiler/CSimpAttr.c @@ -89,6 +89,7 @@ lean_object* l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContex static lean_object* l_Lean_Compiler_CSimp_State_thmNames___default___closed__1; lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_CSimp_initFn____x40_Lean_Compiler_CSimpAttr___hyg_464____closed__15; +lean_object* l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__4(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Compiler_CSimp_replaceConstants___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_CSimp_add___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_ReplaceImpl_replaceUnsafeM_visit___at_Lean_Compiler_CSimp_replaceConstants___spec__7(lean_object*, lean_object*, lean_object*); @@ -135,7 +136,6 @@ static lean_object* l_Lean_Compiler_CSimp_initFn____x40_Lean_Compiler_CSimpAttr_ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Compiler_CSimp_initFn____x40_Lean_Compiler_CSimpAttr___hyg_132____spec__5(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_CSimp_initFn____x40_Lean_Compiler_CSimpAttr___hyg_132_(lean_object*); uint64_t l_Lean_Name_hash___override(lean_object*); -lean_object* l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6189____spec__4(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Compiler_CSimp_initFn____x40_Lean_Compiler_CSimpAttr___hyg_132____spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_CSimp_initFn____x40_Lean_Compiler_CSimpAttr___hyg_464____closed__3; static lean_object* l_Lean_Compiler_CSimp_State_map___default___closed__3; @@ -372,7 +372,7 @@ lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_3 = lean_ctor_get(x_1, 0); x_4 = lean_ctor_get(x_1, 1); x_5 = l_Lean_SMap_switch___at_Lean_Compiler_CSimp_State_switch___spec__1(x_3); -x_6 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6189____spec__4(x_4); +x_6 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__4(x_4); lean_ctor_set(x_1, 1, x_6); lean_ctor_set(x_1, 0, x_5); return x_1; @@ -386,7 +386,7 @@ lean_inc(x_8); lean_inc(x_7); lean_dec(x_1); x_9 = l_Lean_SMap_switch___at_Lean_Compiler_CSimp_State_switch___spec__1(x_7); -x_10 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6189____spec__4(x_8); +x_10 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__4(x_8); x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_9); lean_ctor_set(x_11, 1, x_10); diff --git a/stage0/stdlib/Lean/Compiler/LCNF/AuxDeclCache.c b/stage0/stdlib/Lean/Compiler/LCNF/AuxDeclCache.c index c046278be28e..e893be766f34 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/AuxDeclCache.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/AuxDeclCache.c @@ -26,7 +26,6 @@ lean_object* l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg(lean_object*, lean_object* lean_array_fget(lean_object*, lean_object*); lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_EnvExtension_modifyState___rarg(lean_object*, lean_object*, lean_object*); -uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5345_(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_AuxDeclCache___hyg_9____closed__3; lean_object* l_Lean_EnvExtension_getState___rarg(lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); @@ -42,6 +41,7 @@ static lean_object* l_Lean_Compiler_LCNF_cacheAuxDecl___closed__4; static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_AuxDeclCache___hyg_9____closed__4; lean_object* l_Lean_Compiler_LCNF_cacheAuxDecl___lambda__1(lean_object*, lean_object*, lean_object*); static size_t l_Lean_PersistentHashMap_findAux___at_Lean_Compiler_LCNF_cacheAuxDecl___spec__2___closed__2; +uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5344_(lean_object*, lean_object*); extern lean_object* l_Lean_Compiler_LCNF_instBEqDecl; static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_AuxDeclCache___hyg_9____closed__1; uint64_t l___private_Lean_Compiler_LCNF_DeclHash_0__Lean_Compiler_LCNF_hashDecl____x40_Lean_Compiler_LCNF_DeclHash___hyg_266_(lean_object*); @@ -142,7 +142,7 @@ else lean_object* x_9; uint8_t x_10; x_9 = lean_array_fget(x_1, x_4); lean_inc(x_5); -x_10 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5345_(x_5, x_9); +x_10 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5344_(x_5, x_9); if (x_10 == 0) { lean_object* x_11; lean_object* x_12; @@ -212,7 +212,7 @@ lean_inc(x_11); x_12 = lean_ctor_get(x_10, 1); lean_inc(x_12); lean_dec(x_10); -x_13 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5345_(x_3, x_11); +x_13 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5344_(x_3, x_11); if (x_13 == 0) { lean_object* x_14; @@ -360,7 +360,7 @@ else lean_object* x_17; uint8_t x_18; x_17 = lean_array_fget(x_5, x_2); lean_inc(x_3); -x_18 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5345_(x_3, x_17); +x_18 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5344_(x_3, x_17); if (x_18 == 0) { lean_object* x_19; lean_object* x_20; @@ -458,7 +458,7 @@ x_19 = lean_ctor_get(x_15, 0); x_20 = lean_ctor_get(x_15, 1); lean_inc(x_19); lean_inc(x_4); -x_21 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5345_(x_4, x_19); +x_21 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5344_(x_4, x_19); if (x_21 == 0) { lean_object* x_22; lean_object* x_23; lean_object* x_24; @@ -494,7 +494,7 @@ lean_inc(x_26); lean_dec(x_15); lean_inc(x_26); lean_inc(x_4); -x_28 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5345_(x_4, x_26); +x_28 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5344_(x_4, x_26); if (x_28 == 0) { lean_object* x_29; lean_object* x_30; lean_object* x_31; @@ -617,7 +617,7 @@ if (lean_is_exclusive(x_57)) { } lean_inc(x_60); lean_inc(x_4); -x_63 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5345_(x_4, x_60); +x_63 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5344_(x_4, x_60); if (x_63 == 0) { lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; diff --git a/stage0/stdlib/Lean/Compiler/LCNF/Basic.c b/stage0/stdlib/Lean/Compiler/LCNF/Basic.c index e73ddc209877..e74aa1a6b614 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/Basic.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/Basic.c @@ -53,6 +53,7 @@ LEAN_EXPORT lean_object* l_panic___at___private_Lean_Compiler_LCNF_Basic_0__Lean lean_object* l_Lean_Compiler_LCNF_attachCodeDecls_go(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_instInhabitedFunDeclCore___rarg(lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateFunImp___closed__2; +LEAN_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5344____spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_collectUsedAtExpr(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateAltsImp___spec__1(lean_object*); lean_object* l_Lean_Expr_lit___override(lean_object*); @@ -81,8 +82,6 @@ lean_object* l_Lean_Compiler_LCNF_Arg_toLetValue(lean_object*); lean_object* l_Lean_Compiler_LCNF_instInhabitedAltCore___rarg(lean_object*); static lean_object* l_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateCasesImp(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5345_(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5345____boxed(lean_object*, lean_object*); lean_object* l_Lean_RBNode_insert___at_Lean_NameSet_insert___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_collectArgs___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_instInhabitedAltCore___rarg___closed__1; @@ -182,6 +181,7 @@ static lean_object* l_Lean_Compiler_LCNF_instInhabitedParam___closed__1; static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateUnreachImp___closed__1; lean_object* l_Lean_Compiler_LCNF_Decl_instantiateValueLevelParams_instExpr(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instHashableLetValue; +LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5344_(lean_object*, lean_object*); lean_object* l_Array_mapMono(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_collectArg(lean_object*, lean_object*); uint64_t l_Lean_Expr_hash(lean_object*); @@ -189,6 +189,7 @@ uint8_t l_Lean_BinderInfo_isInstImplicit(uint8_t); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instHashableArg; lean_object* l_Lean_Compiler_LCNF_LitValue_toExpr(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_collectParams___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5344____boxed(lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_markRecDecls(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instBEqParam; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instInhabitedLitValue; @@ -197,7 +198,6 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instBEqCode; uint8_t lean_name_eq(lean_object*, lean_object*); uint8_t l___private_Lean_Compiler_InlineAttrs_0__Lean_Compiler_beqInlineAttributeKind____x40_Lean_Compiler_InlineAttrs___hyg_15_(uint8_t, uint8_t); uint8_t l_Lean_Compiler_LCNF_Decl_inlineIfReduceAttr(lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5345____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instBEqDecl; static lean_object* l_Lean_Compiler_LCNF_instBEqParam___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instInhabitedArg; @@ -214,7 +214,6 @@ lean_object* l_Lean_Compiler_LCNF_CasesCore_getCtorNames(lean_object*); lean_object* l_Lean_Compiler_LCNF_Code_isReturnOf___boxed(lean_object*, lean_object*); lean_object* l_Array_eraseIdx___rarg(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___closed__3; -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5345____spec__1(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Compiler_LCNF_Decl_safe___default; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instHashableLitValue; lean_object* l_Lean_Compiler_LCNF_markRecDecls_visit(lean_object*, lean_object*, lean_object*); @@ -234,11 +233,11 @@ LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_L static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateAltsImp___closed__2; lean_object* l_Lean_Compiler_LCNF_Decl_isCasesOnParam_x3f_go___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instInhabitedCodeDecl; -LEAN_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5345____spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_List_beq___at___private_Lean_Data_OpenDecl_0__Lean_beqOpenDecl____x40_Lean_Data_OpenDecl___hyg_39____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_collectLetValue(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateFVarImp___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_instantiateRangeArgs___spec__1(size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5344____spec__1___boxed(lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_Decl_instantiateValueLevelParams_instFunDecl(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_Decl_instantiateValueLevelParams_instLetValue(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_instHashableArg___closed__1; @@ -345,6 +344,7 @@ uint8_t lean_usize_dec_lt(size_t, size_t); static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateAltImp___closed__2; static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_collectType___closed__2; lean_object* l_Lean_Compiler_LCNF_Code_sizeLe_go(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5344____spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_instInhabitedLitValue___closed__1; static lean_object* l_Lean_Compiler_LCNF_instBEqArg___closed__1; lean_object* lean_nat_add(lean_object*, lean_object*); @@ -370,8 +370,8 @@ lean_object* l_Lean_Compiler_LCNF_Decl_isTemplateLike___boxed(lean_object*, lean lean_object* l_Lean_Compiler_LCNF_hasLocalInst___boxed(lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___spec__1(lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateCasesImp___closed__1; +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5344____spec__1(lean_object*, lean_object*); uint8_t l_Lean_Compiler_LCNF_Decl_isCasesOnParam_x3f_go___lambda__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5345____spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* _init_l_Lean_Compiler_LCNF_instInhabitedParam___closed__1() { _start: { @@ -6675,7 +6675,7 @@ x_1 = l_Lean_Compiler_LCNF_instInhabitedDecl___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5345____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5344____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -6722,7 +6722,7 @@ return x_10; } } } -LEAN_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5345____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 uint8_t l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5344____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; @@ -6764,7 +6764,7 @@ goto _start; } } } -LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5345_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5344_(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; 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; uint8_t x_16; uint8_t x_17; lean_object* x_18; lean_object* x_19; uint8_t x_25; @@ -6879,7 +6879,7 @@ else { lean_object* x_35; uint8_t x_36; x_35 = lean_unsigned_to_nat(0u); -x_36 = l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5345____spec__2(x_6, x_14, lean_box(0), x_6, x_14, x_35); +x_36 = l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5344____spec__2(x_6, x_14, lean_box(0), x_6, x_14, x_35); lean_dec(x_14); lean_dec(x_6); if (x_36 == 0) @@ -6957,7 +6957,7 @@ if (x_9 == 0) if (x_17 == 0) { uint8_t x_20; -x_20 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5345____spec__1(x_10, x_18); +x_20 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5344____spec__1(x_10, x_18); return x_20; } else @@ -6982,27 +6982,27 @@ return x_22; else { uint8_t x_23; -x_23 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5345____spec__1(x_10, x_18); +x_23 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5344____spec__1(x_10, x_18); return x_23; } } } } } -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5345____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5344____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5345____spec__1(x_1, x_2); +x_3 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5344____spec__1(x_1, x_2); x_4 = lean_box(x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5345____spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5344____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: { uint8_t x_7; lean_object* x_8; -x_7 = l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5345____spec__2(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5344____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_2); @@ -7011,11 +7011,11 @@ x_8 = lean_box(x_7); return x_8; } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5345____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5344____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5345_(x_1, x_2); +x_3 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5344_(x_1, x_2); x_4 = lean_box(x_3); return x_4; } @@ -7024,7 +7024,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_instBEqDecl___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5345____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5344____boxed), 2, 0); return x_1; } } diff --git a/stage0/stdlib/Lean/Compiler/LCNF/Simp/JpCases.c b/stage0/stdlib/Lean/Compiler/LCNF/Simp/JpCases.c index fad9e71a5c52..e98f5391ec9a 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/Simp/JpCases.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/Simp/JpCases.c @@ -36,7 +36,6 @@ lean_object* l_Lean_indentD(lean_object*); static lean_object* l_List_forIn_loop___at_Lean_Compiler_LCNF_Simp_simpJpCases_x3f___spec__7___closed__1; lean_object* l_Lean_Compiler_LCNF_Simp_isJpCases_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_mkJpAlt(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_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__8; static lean_object* l_Lean_Compiler_LCNF_Simp_instInhabitedJpCasesInfo___closed__1; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Simp_collectJpCasesInfo_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_panic___at___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_extractJpCases_go___spec__1___closed__2; @@ -53,10 +52,8 @@ uint8_t lean_usize_dec_eq(size_t, size_t); LEAN_EXPORT uint8_t l_Lean_RBNode_any___at_Lean_Compiler_LCNF_Simp_simpJpCases_x3f_visitJp_x3f___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_mkJpAlt_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Simp_simpJpCases_x3f___lambda__1___closed__1; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__2; LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Compiler_LCNF_Simp_simpJpCases_x3f___spec__7(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_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__15; lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Compiler_LCNF_Simp_simpJpCases_x3f_visitJmp_x3f___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_Simp_findCtorName_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -75,16 +72,13 @@ LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Com lean_object* l_Lean_Compiler_LCNF_LCtx_toLocalContext(lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_mkJpAlt_go___closed__2; LEAN_EXPORT lean_object* l_Lean_RBNode_revFold___at_Lean_Compiler_LCNF_Simp_simpJpCases_x3f___spec__4(lean_object*, lean_object*); -lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480_(lean_object*); lean_object* l_Lean_Compiler_LCNF_Simp_simpJpCases_x3f___lambda__1___boxed(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_Compiler_LCNF_Simp_simpJpCases_x3f___spec__7___closed__3; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__5; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_Simp_simpJpCases_x3f_visitJp_x3f___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_ptr_addr(lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_extractJpCases_go___closed__1; size_t lean_usize_of_nat(lean_object*); -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__1; lean_object* lean_st_ref_take(lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_Simp_isJpCases_x3f_go___lambda__1___boxed(lean_object*, lean_object*); @@ -93,17 +87,16 @@ static lean_object* l_Lean_Compiler_LCNF_Simp_collectJpCasesInfo___closed__3; static lean_object* l_Lean_Compiler_LCNF_Simp_simpJpCases_x3f_visitJmp_x3f___closed__1; lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_Internalize_internalizeFunDecl___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_mkJpAlt_go___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_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__13; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__17; 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_Simp_JpCases___hyg_3478____closed__10; LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_LCNF_Simp_simpJpCases_x3f___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_RBNode_any___at_Lean_Compiler_LCNF_Simp_JpCasesInfoMap_isCandidate___spec__1(lean_object*); -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__10; LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_mkJmpArgsAtJp___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_mkJmpArgsAtJp(lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_RBNode_revFold___at_Lean_Compiler_LCNF_Simp_simpJpCases_x3f___spec__4___boxed(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*); -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__7; static lean_object* l_Lean_addTrace___at_Lean_Compiler_LCNF_Simp_simpJpCases_x3f___spec__8___closed__1; lean_object* lean_st_mk_ref(lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_Simp_simpJpCases_x3f_visitJp_x3f___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*); @@ -113,22 +106,24 @@ lean_object* l_Lean_Compiler_LCNF_eraseCode(lean_object*, lean_object*, lean_obj lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_Simp_simpJpCases_x3f_visitJp_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__18; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_mkJpAlt_go___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_JpCasesInfo_ctorNames___default; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__11; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__16; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_mkJpAlt_go___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__6; lean_object* l_Lean_RBNode_insert___at_Lean_FVarIdMap_insert___spec__1___rarg(lean_object*, lean_object*, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_Simp_collectJpCasesInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_extractJpCases(lean_object*); lean_object* l_Lean_Compiler_LCNF_Simp_simpJpCases_x3f_visit___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_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__3; lean_object* l_Lean_Compiler_LCNF_Internalize_internalizeParam(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Simp_collectJpCasesInfo___closed__1; extern lean_object* l_Lean_Compiler_LCNF_instInhabitedArg; lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__19; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__7; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__19; LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_mkJpAlt_go(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* l_Lean_Compiler_LCNF_CasesCore_getCtorNames(lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_mkJpAlt_go___closed__1; @@ -138,19 +133,19 @@ extern lean_object* l_Lean_inheritedTraceOptions; lean_object* l_Lean_MessageData_ofExpr(lean_object*); lean_object* l_Lean_Compiler_LCNF_Simp_simpJpCases_x3f_visitJp_x3f(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_instInhabitedCasesCore(lean_object*); -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__17; lean_object* l_Lean_Compiler_LCNF_attachCodeDecls(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBTree_toList___at_Lean_Compiler_LCNF_Simp_simpJpCases_x3f___spec__3(lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Compiler_LCNF_Simp_simpJpCases_x3f___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_extractJpCases_go(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Simp_simpJpCases_x3f_visitJp_x3f___lambda__1___closed__2; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__4; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__1; static lean_object* l___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_extractJpCases_go___closed__2; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__14; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__3; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Simp_collectJpCasesInfo_go___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*); static lean_object* l___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_mkJpAlt___closed__1; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__4; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__13; extern lean_object* l_Lean_NameSet_empty; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__2; LEAN_EXPORT lean_object* l___private_Init_Data_Array_BasicAux_0__mapMonoMImp___at_Lean_Compiler_LCNF_Simp_simpJpCases_x3f_visit___spec__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_eq(lean_object*, lean_object*); static lean_object* l_Lean_isTracingEnabledFor___at_Lean_Compiler_LCNF_Simp_simpJpCases_x3f___spec__6___closed__1; @@ -162,23 +157,27 @@ LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Compiler_LCNF_Simp_simpJpCa LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_mkJpAlt_go___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*); uint8_t l_Lean_NameSet_contains(lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_Simp_simpJpCases_x3f_visit(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_Simp_JpCases___hyg_3478____closed__14; LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Compiler_LCNF_Simp_simpJpCases_x3f_visitJp_x3f___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_panic_fn(lean_object*, lean_object*); uint8_t l___private_Lean_Compiler_LCNF_DependsOn_0__Lean_Compiler_LCNF_depOn(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Compiler_LCNF_Simp_simpJpCases_x3f___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__12; static lean_object* l_Lean_addTrace___at_Lean_Compiler_LCNF_Simp_simpJpCases_x3f___spec__8___closed__4; static lean_object* l_Lean_Compiler_LCNF_Simp_simpJpCases_x3f___lambda__1___closed__7; lean_object* lean_nat_sub(lean_object*, lean_object*); extern lean_object* l_Lean_Compiler_LCNF_instInhabitedParam; static lean_object* l_panic___at___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_extractJpCases_go___spec__1___closed__1; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__9; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_mkJpAlt_go___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478_(lean_object*); +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__8; lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); 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*); lean_object* l_Lean_Compiler_LCNF_Simp_JpCasesInfoMap_isCandidate___boxed(lean_object*); static lean_object* l_panic___at___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_extractJpCases_go___spec__1___closed__3; lean_object* l_List_reverse___rarg(lean_object*); +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__6; lean_object* l_Lean_mkHashMapImp___rarg(lean_object*); uint8_t l_Lean_Name_quickCmp(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_BasicAux_0__mapMonoMImp_go___at_Lean_Compiler_LCNF_Simp_simpJpCases_x3f_visit___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -186,7 +185,6 @@ lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateAl size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_LCNF_Simp_simpJpCases_x3f___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_addTrace___at_Lean_Compiler_LCNF_Simp_simpJpCases_x3f___spec__8___closed__3; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__12; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_mkJpAlt_go___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_object*, lean_object*, lean_object*); uint8_t l_Lean_Compiler_LCNF_CodeDecl_dependsOn(lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); @@ -200,19 +198,21 @@ static lean_object* l_Lean_Compiler_LCNF_Simp_simpJpCases_x3f___lambda__1___clos lean_object* lean_array_get_size(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_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__16; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__5; lean_object* l_Lean_Compiler_LCNF_Simp_isJpCases_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Simp_collectJpCasesInfo___closed__4; lean_object* l_Lean_Compiler_LCNF_mkAuxFunDecl(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_Simp_JpCases___hyg_3480____closed__18; lean_object* l_Lean_Compiler_LCNF_Simp_collectJpCasesInfo_go___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_Compiler_LCNF_Simp_instInhabitedJpCasesInfo; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__15; static lean_object* l_Lean_Compiler_LCNF_Simp_collectJpCasesInfo___closed__2; static lean_object* l_Lean_addTrace___at_Lean_Compiler_LCNF_Simp_simpJpCases_x3f___spec__8___closed__2; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__9; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); lean_object* l_Lean_MessageData_ofName(lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Simp_JpCases_0__Lean_Compiler_LCNF_Simp_extractJpCases_go___closed__4; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__11; static lean_object* l_List_forIn_loop___at_Lean_Compiler_LCNF_Simp_simpJpCases_x3f___spec__7___closed__4; lean_object* l_Lean_Compiler_LCNF_Simp_isJpCases_x3f_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Compiler_LCNF_Simp_collectJpCasesInfo_go___spec__1___boxed(lean_object*, lean_object*); @@ -7969,7 +7969,7 @@ lean_dec(x_3); return x_9; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__1() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__1() { _start: { lean_object* x_1; @@ -7977,27 +7977,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_Simp_JpCases___hyg_3480____closed__2() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__2() { _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_Simp_JpCases___hyg_3480____closed__1; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____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_Simp_JpCases___hyg_3480____closed__3() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__2; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__2; x_2 = l_Lean_Compiler_LCNF_Simp_simpJpCases_x3f___lambda__1___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_Simp_JpCases___hyg_3480____closed__4() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__4() { _start: { lean_object* x_1; @@ -8005,17 +8005,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_Simp_JpCases___hyg_3480____closed__5() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____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_Simp_JpCases___hyg_3480____closed__3; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__4; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__3; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__4; 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_Simp_JpCases___hyg_3480____closed__6() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__6() { _start: { lean_object* x_1; @@ -8023,17 +8023,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_Simp_JpCases___hyg_3480____closed__7() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____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_Simp_JpCases___hyg_3480____closed__5; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__6; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__5; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____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_Simp_JpCases___hyg_3480____closed__8() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__8() { _start: { lean_object* x_1; @@ -8041,47 +8041,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_Simp_JpCases___hyg_3480____closed__9() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____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_Simp_JpCases___hyg_3480____closed__7; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__8; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__7; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____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_Simp_JpCases___hyg_3480____closed__10() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__9; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__1; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__9; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____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_Simp_JpCases___hyg_3480____closed__11() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____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_Simp_JpCases___hyg_3480____closed__10; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__10; x_2 = l_Lean_Compiler_LCNF_Simp_simpJpCases_x3f___lambda__1___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_Simp_JpCases___hyg_3480____closed__12() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____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_Simp_JpCases___hyg_3480____closed__11; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__4; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__11; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__4; 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_Simp_JpCases___hyg_3480____closed__13() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__13() { _start: { lean_object* x_1; @@ -8089,17 +8089,17 @@ x_1 = lean_mk_string_from_bytes("Simp", 4); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__14() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____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_Simp_JpCases___hyg_3480____closed__12; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__13; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__12; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__13; 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_Simp_JpCases___hyg_3480____closed__15() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__15() { _start: { lean_object* x_1; @@ -8107,17 +8107,17 @@ x_1 = lean_mk_string_from_bytes("JpCases", 7); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__16() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____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_Simp_JpCases___hyg_3480____closed__14; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__15; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__14; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____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_Simp_JpCases___hyg_3480____closed__17() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__17() { _start: { lean_object* x_1; @@ -8125,33 +8125,33 @@ 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_Simp_JpCases___hyg_3480____closed__18() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____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_Simp_JpCases___hyg_3480____closed__16; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__17; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__16; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____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_Simp_JpCases___hyg_3480____closed__19() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____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_Simp_JpCases___hyg_3480____closed__18; -x_2 = lean_unsigned_to_nat(3480u); +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__18; +x_2 = lean_unsigned_to_nat(3478u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480_(lean_object* x_1) { +lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478_(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_Simp_simpJpCases_x3f___lambda__1___closed__4; x_3 = 0; -x_4 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__19; +x_4 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__19; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -8257,45 +8257,45 @@ l_Lean_Compiler_LCNF_Simp_simpJpCases_x3f___lambda__1___closed__6 = _init_l_Lean lean_mark_persistent(l_Lean_Compiler_LCNF_Simp_simpJpCases_x3f___lambda__1___closed__6); l_Lean_Compiler_LCNF_Simp_simpJpCases_x3f___lambda__1___closed__7 = _init_l_Lean_Compiler_LCNF_Simp_simpJpCases_x3f___lambda__1___closed__7(); lean_mark_persistent(l_Lean_Compiler_LCNF_Simp_simpJpCases_x3f___lambda__1___closed__7); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__1 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__1(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__1); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__2 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__2(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__2); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__3 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__3(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__3); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__4 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__4(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__4); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__5 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__5(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__5); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__6 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__6(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__6); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__7 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__7(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__7); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__8 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__8(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__8); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__9 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__9(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__9); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__10 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__10(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__10); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__11 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__11(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__11); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__12 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__12(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__12); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__13 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__13(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__13); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__14 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__14(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__14); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__15 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__15(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__15); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__16 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__16(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__16); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__17 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__17(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__17); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__18 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__18(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__18); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__19 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__19(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480____closed__19); -if (builtin) {res = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3480_(lean_io_mk_world()); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__1 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__1(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__1); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__2 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__2(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__2); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__3 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__3(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__3); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__4 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__4(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__4); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__5 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__5(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__5); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__6 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__6(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__6); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__7 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__7(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__7); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__8 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__8(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__8); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__9 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__9(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__9); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__10 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__10(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__10); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__11 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__11(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__11); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__12 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__12(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__12); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__13 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__13(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__13); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__14 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__14(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__14); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__15 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__15(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__15); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__16 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__16(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__16); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__17 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__17(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__17); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__18 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__18(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__18); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__19 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__19(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478____closed__19); +if (builtin) {res = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Simp_JpCases___hyg_3478_(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/Testing.c b/stage0/stdlib/Lean/Compiler/LCNF/Testing.c index 164040e8eb88..225ca523a27a 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/Testing.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/Testing.c @@ -53,7 +53,6 @@ lean_object* lean_array_fget(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Testing_assertSize___lambda__1___closed__3; lean_object* l_Lean_Compiler_LCNF_Testing_assert___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Testing_assertIsAtFixPoint___lambda__1___closed__3; -uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5345_(lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_Code_containsConst_goLetValue___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Testing_assertIsAtFixPoint___lambda__1___closed__2; static lean_object* l___private_Lean_Compiler_LCNF_Testing_0__Lean_Compiler_LCNF_Testing_assertAfterTest___elambda__1___closed__6; @@ -112,6 +111,7 @@ lean_object* l_Lean_Compiler_LCNF_AltCore_getCode(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Testing_0__Lean_Compiler_LCNF_Testing_throwFixPointError(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_Testing_getDecls___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_Testing_assert(uint8_t, 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_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5344_(lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_Testing_TestM_run___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Testing_assertIsAtFixPoint___lambda__1___closed__1; uint8_t lean_name_eq(lean_object*, lean_object*); @@ -2894,7 +2894,7 @@ else lean_object* x_10; lean_object* x_11; uint8_t x_12; x_10 = lean_array_fget(x_4, x_6); x_11 = lean_array_fget(x_5, x_6); -x_12 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5345_(x_10, x_11); +x_12 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5344_(x_10, x_11); if (x_12 == 0) { uint8_t x_13; diff --git a/stage0/stdlib/Lean/CoreM.c b/stage0/stdlib/Lean/CoreM.c index 4c227fb18579..586df1d2b2c8 100644 --- a/stage0/stdlib/Lean/CoreM.c +++ b/stage0/stdlib/Lean/CoreM.c @@ -13880,7 +13880,7 @@ 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) +if (lean_obj_tag(x_12) == 12) { lean_object* x_13; lean_object* x_14; x_13 = lean_ctor_get(x_12, 0); @@ -14097,7 +14097,7 @@ 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) +if (lean_obj_tag(x_12) == 12) { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; x_13 = lean_ctor_get(x_12, 0); diff --git a/stage0/stdlib/Lean/Data/Json/Parser.c b/stage0/stdlib/Lean/Data/Json/Parser.c index 61fba5db3d35..27562bc511d0 100644 --- a/stage0/stdlib/Lean/Data/Json/Parser.c +++ b/stage0/stdlib/Lean/Data/Json/Parser.c @@ -48,8 +48,8 @@ lean_object* lean_string_push(lean_object*, uint32_t); lean_object* l_Lean_Parsec_pstring(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_Parser_natMaybeZero(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_Parser_escapedChar___boxed__const__7; -static lean_object* l_Lean_Json_Parser_num___lambda__3___closed__3; uint32_t l_String_Iterator_curr(lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____at_Lean_Json_Parser_num___spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Json_Parser_num___lambda__5___closed__2; LEAN_EXPORT lean_object* l_Lean_Json_Parser_escapedChar___boxed__const__3; static lean_object* l_Lean_Json_Parser_anyCore___closed__6; @@ -69,12 +69,12 @@ static lean_object* l_Lean_Json_Parser_objectCore___closed__2; static lean_object* l_Lean_Json_Parser_anyCore___closed__3; static lean_object* l_Lean_Json_Parser_num___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Json_Parser_escapedChar___boxed__const__5; -LEAN_EXPORT lean_object* l_Lean_Json_Parser_num___lambda__3___closed__3___boxed__const__1; static lean_object* l_Lean_Json_parse___closed__2; static lean_object* l_Lean_Json_Parser_lookahead___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Json_Parser_natNonZero(lean_object*); static lean_object* l_Lean_Json_Parser_num___closed__3; extern lean_object* l_System_Platform_numBits; +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____at_Lean_Json_Parser_num___spec__1(lean_object*, lean_object*); extern lean_object* l_Lean_Parsec_expectedEndOfInput; static lean_object* l_Lean_Json_Parser_num___lambda__2___closed__3; extern lean_object* l_Std_Format_defWidth; @@ -87,13 +87,13 @@ static lean_object* l_Lean_Json_Parser_objectCore___closed__1; LEAN_EXPORT lean_object* l_Lean_Json_Parser_strCore(lean_object*, lean_object*); lean_object* lean_int_mul(lean_object*, lean_object*); lean_object* lean_nat_pow(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_Parser_num___lambda__3___closed__2___boxed__const__1; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); uint8_t lean_uint32_dec_eq(uint32_t, uint32_t); static lean_object* l_Lean_Json_Parser_escapedChar___closed__1; static lean_object* l_Lean_Json_parse___closed__1; static lean_object* l_Lean_Json_Parser_anyCore___closed__7; LEAN_EXPORT lean_object* l_Lean_Json_Parser_num(lean_object*); -lean_object* l_UInt32_decEq___boxed(lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); static lean_object* l_Lean_Json_Parser_natNonZero___closed__2; LEAN_EXPORT lean_object* l_Lean_Json_Parser_num___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -112,7 +112,6 @@ LEAN_EXPORT lean_object* l_Lean_Json_Parser_arrayCore(lean_object*, lean_object* LEAN_EXPORT lean_object* l_Lean_Json_parse(lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); static lean_object* l_Lean_Json_Parser_num___closed__1; -lean_object* l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____at_Option_instDecidableEqOption___spec__1___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_RBNode_insert___at_Lean_Json_mkObj___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_Parser_escapedChar___boxed__const__8; static lean_object* l_Lean_Json_Parser_natNonZero___closed__1; @@ -1698,6 +1697,53 @@ return x_22; } } } +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____at_Lean_Json_Parser_num___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_3; +x_3 = 1; +return x_3; +} +else +{ +uint8_t x_4; +lean_dec(x_2); +x_4 = 0; +return x_4; +} +} +else +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_5; +lean_dec(x_1); +x_5 = 0; +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; uint32_t x_8; uint32_t x_9; uint8_t x_10; +x_6 = lean_ctor_get(x_1, 0); +lean_inc(x_6); +lean_dec(x_1); +x_7 = lean_ctor_get(x_2, 0); +lean_inc(x_7); +lean_dec(x_2); +x_8 = lean_unbox_uint32(x_6); +lean_dec(x_6); +x_9 = lean_unbox_uint32(x_7); +lean_dec(x_7); +x_10 = lean_uint32_dec_eq(x_8, x_9); +return x_10; +} +} +} +} LEAN_EXPORT lean_object* l_Lean_Json_Parser_num___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { @@ -1887,15 +1933,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Json_Parser_num___lambda__3___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_UInt32_decEq___boxed), 2, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_Json_Parser_num___lambda__3___closed__3___boxed__const__1() { +static lean_object* _init_l_Lean_Json_Parser_num___lambda__3___closed__2___boxed__const__1() { _start: { uint32_t x_1; lean_object* x_2; @@ -1904,11 +1942,11 @@ x_2 = lean_box_uint32(x_1); return x_2; } } -static lean_object* _init_l_Lean_Json_Parser_num___lambda__3___closed__3() { +static lean_object* _init_l_Lean_Json_Parser_num___lambda__3___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Json_Parser_num___lambda__3___closed__3___boxed__const__1; +x_1 = l_Lean_Json_Parser_num___lambda__3___closed__2___boxed__const__1; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; @@ -1917,26 +1955,26 @@ return x_2; LEAN_EXPORT lean_object* l_Lean_Json_Parser_num___lambda__3(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_3; uint32_t x_4; lean_object* x_39; lean_object* x_40; uint8_t x_60; -x_60 = l_String_Iterator_hasNext(x_2); -if (x_60 == 0) +lean_object* x_3; uint32_t x_4; lean_object* x_39; lean_object* x_40; uint8_t x_57; +x_57 = l_String_Iterator_hasNext(x_2); +if (x_57 == 0) { -lean_object* x_61; -x_61 = lean_box(0); +lean_object* x_58; +x_58 = lean_box(0); x_39 = x_2; -x_40 = x_61; -goto block_59; +x_40 = x_58; +goto block_56; } else { -uint32_t x_62; lean_object* x_63; lean_object* x_64; -x_62 = l_String_Iterator_curr(x_2); -x_63 = lean_box_uint32(x_62); -x_64 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_64, 0, x_63); +uint32_t x_59; lean_object* x_60; lean_object* x_61; +x_59 = l_String_Iterator_curr(x_2); +x_60 = lean_box_uint32(x_59); +x_61 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_61, 0, x_60); x_39 = x_2; -x_40 = x_64; -goto block_59; +x_40 = x_61; +goto block_56; } block_38: { @@ -2051,77 +2089,72 @@ return x_37; } } } -block_59: +block_56: { -lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; -x_41 = l_Lean_Json_Parser_num___lambda__3___closed__2; -x_42 = l_Lean_Json_Parser_num___lambda__3___closed__1; +lean_object* x_41; uint8_t x_42; +x_41 = l_Lean_Json_Parser_num___lambda__3___closed__1; lean_inc(x_40); -x_43 = l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____at_Option_instDecidableEqOption___spec__1___rarg(x_41, x_40, x_42); -x_44 = lean_unbox(x_43); -lean_dec(x_43); -if (x_44 == 0) +x_42 = l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____at_Lean_Json_Parser_num___spec__1(x_40, x_41); +if (x_42 == 0) { -lean_object* x_45; lean_object* x_46; uint8_t x_47; -x_45 = l_Lean_Json_Parser_num___lambda__3___closed__3; -x_46 = l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____at_Option_instDecidableEqOption___spec__1___rarg(x_41, x_40, x_45); -x_47 = lean_unbox(x_46); -lean_dec(x_46); -if (x_47 == 0) +lean_object* x_43; uint8_t x_44; +x_43 = l_Lean_Json_Parser_num___lambda__3___closed__2; +x_44 = l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____at_Lean_Json_Parser_num___spec__1(x_40, x_43); +if (x_44 == 0) { -lean_object* x_48; -x_48 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_48, 0, x_39); -lean_ctor_set(x_48, 1, x_1); -return x_48; +lean_object* x_45; +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_39); +lean_ctor_set(x_45, 1, x_1); +return x_45; } else { -lean_object* x_49; uint8_t x_50; -x_49 = l_String_Iterator_next(x_39); -x_50 = l_String_Iterator_hasNext(x_49); -if (x_50 == 0) +lean_object* x_46; uint8_t x_47; +x_46 = l_String_Iterator_next(x_39); +x_47 = l_String_Iterator_hasNext(x_46); +if (x_47 == 0) { -lean_object* x_51; lean_object* x_52; +lean_object* x_48; lean_object* x_49; lean_dec(x_1); -x_51 = l_Lean_Parsec_unexpectedEndOfInput; -x_52 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_52, 0, x_49); -lean_ctor_set(x_52, 1, x_51); -return x_52; +x_48 = l_Lean_Parsec_unexpectedEndOfInput; +x_49 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_49, 0, x_46); +lean_ctor_set(x_49, 1, x_48); +return x_49; } else { -uint32_t x_53; -x_53 = l_String_Iterator_curr(x_49); -x_3 = x_49; -x_4 = x_53; +uint32_t x_50; +x_50 = l_String_Iterator_curr(x_46); +x_3 = x_46; +x_4 = x_50; goto block_38; } } } else { -lean_object* x_54; uint8_t x_55; +lean_object* x_51; uint8_t x_52; lean_dec(x_40); -x_54 = l_String_Iterator_next(x_39); -x_55 = l_String_Iterator_hasNext(x_54); -if (x_55 == 0) +x_51 = l_String_Iterator_next(x_39); +x_52 = l_String_Iterator_hasNext(x_51); +if (x_52 == 0) { -lean_object* x_56; lean_object* x_57; +lean_object* x_53; lean_object* x_54; lean_dec(x_1); -x_56 = l_Lean_Parsec_unexpectedEndOfInput; -x_57 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_57, 0, x_54); -lean_ctor_set(x_57, 1, x_56); -return x_57; +x_53 = l_Lean_Parsec_unexpectedEndOfInput; +x_54 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_54, 0, x_51); +lean_ctor_set(x_54, 1, x_53); +return x_54; } else { -uint32_t x_58; -x_58 = l_String_Iterator_curr(x_54); -x_3 = x_54; -x_4 = x_58; +uint32_t x_55; +x_55 = l_String_Iterator_curr(x_51); +x_3 = x_51; +x_4 = x_55; goto block_38; } } @@ -2190,174 +2223,171 @@ return x_1; LEAN_EXPORT lean_object* l_Lean_Json_Parser_num___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_4; lean_object* x_5; uint8_t x_52; -x_52 = l_String_Iterator_hasNext(x_3); -if (x_52 == 0) +lean_object* x_4; lean_object* x_5; uint8_t x_50; +x_50 = l_String_Iterator_hasNext(x_3); +if (x_50 == 0) { -lean_object* x_53; -x_53 = lean_box(0); +lean_object* x_51; +x_51 = lean_box(0); x_4 = x_3; -x_5 = x_53; -goto block_51; +x_5 = x_51; +goto block_49; } else { -uint32_t x_54; lean_object* x_55; lean_object* x_56; -x_54 = l_String_Iterator_curr(x_3); -x_55 = lean_box_uint32(x_54); -x_56 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_56, 0, x_55); +uint32_t x_52; lean_object* x_53; lean_object* x_54; +x_52 = l_String_Iterator_curr(x_3); +x_53 = lean_box_uint32(x_52); +x_54 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_54, 0, x_53); x_4 = x_3; -x_5 = x_56; -goto block_51; +x_5 = x_54; +goto block_49; } -block_51: +block_49: { -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; +lean_object* x_6; lean_object* x_7; uint8_t x_8; x_6 = l_Lean_Json_Parser_num___lambda__5___closed__1; -x_7 = l_Lean_Json_Parser_num___lambda__3___closed__2; -x_8 = l_Lean_Json_Parser_num___lambda__5___closed__2; -x_9 = l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____at_Option_instDecidableEqOption___spec__1___rarg(x_7, x_5, x_8); -x_10 = lean_unbox(x_9); -lean_dec(x_9); -if (x_10 == 0) +x_7 = l_Lean_Json_Parser_num___lambda__5___closed__2; +x_8 = l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____at_Lean_Json_Parser_num___spec__1(x_5, x_7); +if (x_8 == 0) { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_11 = lean_nat_to_int(x_2); -x_12 = lean_int_mul(x_1, x_11); -lean_dec(x_11); +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_9 = lean_nat_to_int(x_2); +x_10 = lean_int_mul(x_1, x_9); +lean_dec(x_9); lean_dec(x_1); -x_13 = lean_unsigned_to_nat(0u); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_12); -lean_ctor_set(x_14, 1, x_13); -x_15 = lean_apply_2(x_6, x_14, x_4); -return x_15; +x_11 = lean_unsigned_to_nat(0u); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_10); +lean_ctor_set(x_12, 1, x_11); +x_13 = lean_apply_2(x_6, x_12, x_4); +return x_13; } else { -lean_object* x_16; uint8_t x_17; -x_16 = l_String_Iterator_next(x_4); -x_17 = l_String_Iterator_hasNext(x_16); -if (x_17 == 0) +lean_object* x_14; uint8_t x_15; +x_14 = l_String_Iterator_next(x_4); +x_15 = l_String_Iterator_hasNext(x_14); +if (x_15 == 0) { -lean_object* x_18; lean_object* x_19; +lean_object* x_16; lean_object* x_17; lean_dec(x_2); lean_dec(x_1); -x_18 = l_Lean_Parsec_unexpectedEndOfInput; -x_19 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_19, 0, x_16); -lean_ctor_set(x_19, 1, x_18); -return x_19; +x_16 = l_Lean_Parsec_unexpectedEndOfInput; +x_17 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_17, 0, x_14); +lean_ctor_set(x_17, 1, x_16); +return x_17; } else { -uint32_t x_20; uint32_t x_21; uint8_t x_22; -x_20 = l_String_Iterator_curr(x_16); -x_21 = 48; -x_22 = lean_uint32_dec_le(x_21, x_20); -if (x_22 == 0) +uint32_t x_18; uint32_t x_19; uint8_t x_20; +x_18 = l_String_Iterator_curr(x_14); +x_19 = 48; +x_20 = lean_uint32_dec_le(x_19, x_18); +if (x_20 == 0) { -lean_object* x_23; lean_object* x_24; +lean_object* x_21; lean_object* x_22; lean_dec(x_2); lean_dec(x_1); -x_23 = l_Lean_Json_Parser_natNumDigits___closed__2; -x_24 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_24, 0, x_16); -lean_ctor_set(x_24, 1, x_23); -return x_24; +x_21 = l_Lean_Json_Parser_natNumDigits___closed__2; +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_14); +lean_ctor_set(x_22, 1, x_21); +return x_22; } else { -uint32_t x_25; uint8_t x_26; -x_25 = 57; -x_26 = lean_uint32_dec_le(x_20, x_25); -if (x_26 == 0) +uint32_t x_23; uint8_t x_24; +x_23 = 57; +x_24 = lean_uint32_dec_le(x_18, x_23); +if (x_24 == 0) { -lean_object* x_27; lean_object* x_28; +lean_object* x_25; lean_object* x_26; lean_dec(x_2); lean_dec(x_1); -x_27 = l_Lean_Json_Parser_natNumDigits___closed__2; -x_28 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_28, 0, x_16); -lean_ctor_set(x_28, 1, x_27); -return x_28; +x_25 = l_Lean_Json_Parser_natNumDigits___closed__2; +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_14); +lean_ctor_set(x_26, 1, x_25); +return x_26; } else { -lean_object* x_29; lean_object* x_30; uint8_t x_31; -x_29 = lean_unsigned_to_nat(0u); -x_30 = l_Lean_Json_Parser_natCore(x_29, x_29, x_16); -x_31 = !lean_is_exclusive(x_30); -if (x_31 == 0) +lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_27 = lean_unsigned_to_nat(0u); +x_28 = l_Lean_Json_Parser_natCore(x_27, x_27, x_14); +x_29 = !lean_is_exclusive(x_28); +if (x_29 == 0) { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_32 = lean_ctor_get(x_30, 1); -x_33 = lean_ctor_get(x_30, 0); -x_34 = lean_ctor_get(x_32, 0); -lean_inc(x_34); -x_35 = lean_ctor_get(x_32, 1); -lean_inc(x_35); -lean_dec(x_32); -x_36 = l_Lean_Json_Parser_num___lambda__2___closed__3; -x_37 = lean_nat_dec_lt(x_36, x_35); -if (x_37 == 0) +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_30 = lean_ctor_get(x_28, 1); +x_31 = lean_ctor_get(x_28, 0); +x_32 = lean_ctor_get(x_30, 0); +lean_inc(x_32); +x_33 = lean_ctor_get(x_30, 1); +lean_inc(x_33); +lean_dec(x_30); +x_34 = l_Lean_Json_Parser_num___lambda__2___closed__3; +x_35 = lean_nat_dec_lt(x_34, x_33); +if (x_35 == 0) { -lean_object* x_38; lean_object* x_39; -lean_free_object(x_30); -x_38 = lean_box(0); -x_39 = l_Lean_Json_Parser_num___lambda__4(x_2, x_35, x_34, x_1, x_6, x_38, x_33); +lean_object* x_36; lean_object* x_37; +lean_free_object(x_28); +x_36 = lean_box(0); +x_37 = l_Lean_Json_Parser_num___lambda__4(x_2, x_33, x_32, x_1, x_6, x_36, x_31); lean_dec(x_1); -return x_39; +return x_37; } else { -lean_object* x_40; -lean_dec(x_35); -lean_dec(x_34); +lean_object* x_38; +lean_dec(x_33); +lean_dec(x_32); lean_dec(x_2); lean_dec(x_1); -x_40 = l_Lean_Json_Parser_num___lambda__5___closed__3; -lean_ctor_set_tag(x_30, 1); -lean_ctor_set(x_30, 1, x_40); -return x_30; +x_38 = l_Lean_Json_Parser_num___lambda__5___closed__3; +lean_ctor_set_tag(x_28, 1); +lean_ctor_set(x_28, 1, x_38); +return x_28; } } else { -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; -x_41 = lean_ctor_get(x_30, 1); -x_42 = lean_ctor_get(x_30, 0); +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_39 = lean_ctor_get(x_28, 1); +x_40 = lean_ctor_get(x_28, 0); +lean_inc(x_39); +lean_inc(x_40); +lean_dec(x_28); +x_41 = lean_ctor_get(x_39, 0); lean_inc(x_41); +x_42 = lean_ctor_get(x_39, 1); lean_inc(x_42); -lean_dec(x_30); -x_43 = lean_ctor_get(x_41, 0); -lean_inc(x_43); -x_44 = lean_ctor_get(x_41, 1); -lean_inc(x_44); -lean_dec(x_41); -x_45 = l_Lean_Json_Parser_num___lambda__2___closed__3; -x_46 = lean_nat_dec_lt(x_45, x_44); -if (x_46 == 0) +lean_dec(x_39); +x_43 = l_Lean_Json_Parser_num___lambda__2___closed__3; +x_44 = lean_nat_dec_lt(x_43, x_42); +if (x_44 == 0) { -lean_object* x_47; lean_object* x_48; -x_47 = lean_box(0); -x_48 = l_Lean_Json_Parser_num___lambda__4(x_2, x_44, x_43, x_1, x_6, x_47, x_42); +lean_object* x_45; lean_object* x_46; +x_45 = lean_box(0); +x_46 = l_Lean_Json_Parser_num___lambda__4(x_2, x_42, x_41, x_1, x_6, x_45, x_40); lean_dec(x_1); -return x_48; +return x_46; } else { -lean_object* x_49; lean_object* x_50; -lean_dec(x_44); -lean_dec(x_43); +lean_object* x_47; lean_object* x_48; +lean_dec(x_42); +lean_dec(x_41); lean_dec(x_2); lean_dec(x_1); -x_49 = l_Lean_Json_Parser_num___lambda__5___closed__3; -x_50 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_50, 0, x_42); -lean_ctor_set(x_50, 1, x_49); -return x_50; +x_47 = l_Lean_Json_Parser_num___lambda__5___closed__3; +x_48 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_48, 0, x_40); +lean_ctor_set(x_48, 1, x_47); +return x_48; } } } @@ -2512,6 +2542,15 @@ return x_13; } } } +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____at_Lean_Json_Parser_num___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____at_Lean_Json_Parser_num___spec__1(x_1, x_2); +x_4 = lean_box(x_3); +return x_4; +} +} LEAN_EXPORT lean_object* l_Lean_Json_Parser_num___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { @@ -4234,12 +4273,10 @@ l_Lean_Json_Parser_num___lambda__3___closed__1___boxed__const__1 = _init_l_Lean_ lean_mark_persistent(l_Lean_Json_Parser_num___lambda__3___closed__1___boxed__const__1); l_Lean_Json_Parser_num___lambda__3___closed__1 = _init_l_Lean_Json_Parser_num___lambda__3___closed__1(); lean_mark_persistent(l_Lean_Json_Parser_num___lambda__3___closed__1); +l_Lean_Json_Parser_num___lambda__3___closed__2___boxed__const__1 = _init_l_Lean_Json_Parser_num___lambda__3___closed__2___boxed__const__1(); +lean_mark_persistent(l_Lean_Json_Parser_num___lambda__3___closed__2___boxed__const__1); l_Lean_Json_Parser_num___lambda__3___closed__2 = _init_l_Lean_Json_Parser_num___lambda__3___closed__2(); lean_mark_persistent(l_Lean_Json_Parser_num___lambda__3___closed__2); -l_Lean_Json_Parser_num___lambda__3___closed__3___boxed__const__1 = _init_l_Lean_Json_Parser_num___lambda__3___closed__3___boxed__const__1(); -lean_mark_persistent(l_Lean_Json_Parser_num___lambda__3___closed__3___boxed__const__1); -l_Lean_Json_Parser_num___lambda__3___closed__3 = _init_l_Lean_Json_Parser_num___lambda__3___closed__3(); -lean_mark_persistent(l_Lean_Json_Parser_num___lambda__3___closed__3); l_Lean_Json_Parser_num___lambda__5___closed__1 = _init_l_Lean_Json_Parser_num___lambda__5___closed__1(); lean_mark_persistent(l_Lean_Json_Parser_num___lambda__5___closed__1); l_Lean_Json_Parser_num___lambda__5___closed__2___boxed__const__1 = _init_l_Lean_Json_Parser_num___lambda__5___closed__2___boxed__const__1(); diff --git a/stage0/stdlib/Lean/Data/Position.c b/stage0/stdlib/Lean/Data/Position.c index fc92b0acb117..55a53b2be112 100644 --- a/stage0/stdlib/Lean/Data/Position.c +++ b/stage0/stdlib/Lean/Data/Position.c @@ -101,12 +101,12 @@ LEAN_EXPORT lean_object* l_Lean_instInhabitedPosition; static lean_object* l_Lean_Position_instToExprPosition___lambda__1___closed__2; lean_object* lean_nat_sub(lean_object*, lean_object*); static lean_object* l_Lean_FileMap_ofString___closed__3; -lean_object* l_Nat_decEq___boxed(lean_object*, lean_object*); lean_object* l_Array_back___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Position_instToExprPosition___lambda__1___closed__6; static lean_object* l___private_Lean_Data_Position_0__Lean_reprPosition____x40_Lean_Data_Position___hyg_176____closed__1; static lean_object* l_Lean_Position_instToFormatPosition___closed__1; LEAN_EXPORT lean_object* l_Lean_Position_lt(lean_object*, lean_object*); +lean_object* l_instDecidableEqNat___boxed(lean_object*, lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Position_0__Lean_reprPosition____x40_Lean_Data_Position___hyg_176_(lean_object*, lean_object*); @@ -460,7 +460,7 @@ static lean_object* _init_l_Lean_Position_lt___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Nat_decEq___boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l_instDecidableEqNat___boxed), 2, 0); return x_1; } } diff --git a/stage0/stdlib/Lean/Elab/Command.c b/stage0/stdlib/Lean/Elab/Command.c index 00c6862e616c..87baa0ae135d 100644 --- a/stage0/stdlib/Lean/Elab/Command.c +++ b/stage0/stdlib/Lean/Elab/Command.c @@ -391,6 +391,7 @@ 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_object* l_instDecidableEqPos___boxed(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; @@ -793,7 +794,6 @@ LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at_Lean_Elab_Command_ex lean_object* lean_erase_macro_scopes(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_expandDeclId___spec__18(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_548____closed__13; -lean_object* l_Nat_decEq___boxed(lean_object*, lean_object*); 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*); @@ -4896,7 +4896,7 @@ static lean_object* _init_l___private_Lean_Elab_Command_0__Lean_Elab_Command_add _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Nat_decEq___boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l_instDecidableEqPos___boxed), 2, 0); return x_1; } } diff --git a/stage0/stdlib/Lean/Elab/GuardMsgs.c b/stage0/stdlib/Lean/Elab/GuardMsgs.c index 4b0d9d054c5d..49858809e5d0 100644 --- a/stage0/stdlib/Lean/Elab/GuardMsgs.c +++ b/stage0/stdlib/Lean/Elab/GuardMsgs.c @@ -26,6 +26,7 @@ lean_object* l_Lean_logAt___at_Lean_Elab_Command_elabCommand___spec__4(lean_obje static lean_object* l___private_Lean_Elab_GuardMsgs_0__Lean_Elab_Tactic_GuardMsgs_messageToStringWithoutPos___lambda__3___closed__3; LEAN_EXPORT uint8_t l_Lean_Elab_Tactic_GuardMsgs_parseGuardMsgsSpec___elambda__1(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs(lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare____x40_Lean_Elab_GuardMsgs___hyg_2285____closed__1; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_GuardMsgs_parseGuardMsgsSpec___spec__6___closed__3; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___spec__4(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_GuardMsgs_parseGuardMsgsSpec___spec__6___lambda__4(lean_object*, lean_object*); @@ -49,6 +50,7 @@ lean_object* l_Lean_Syntax_getArgs(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_replaceRef(lean_object*, lean_object*); lean_object* l_String_trim(lean_object*); +static lean_object* l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace___closed__5; lean_object* l_Lean_Syntax_getPos_x3f(lean_object*, uint8_t); lean_object* lean_array_fget(lean_object*, lean_object*); lean_object* l_Lean_Lsp_WorkspaceEdit_ofTextEdit(lean_object*, lean_object*); @@ -60,6 +62,8 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_GuardMsgs_elabG lean_object* l_Lean_addBuiltinDeclarationRanges(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___closed__5; static lean_object* l___private_Lean_Elab_GuardMsgs_0__Lean_Elab_Tactic_GuardMsgs_messageToStringWithoutPos___lambda__3___closed__4; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_GuardMsgs_removeTrailingWhitespaceMarker___boxed(lean_object*); static lean_object* l_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_GuardMsgs_parseGuardMsgsSpec___lambda__2___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_GuardMsgs_0__Lean_Elab_Tactic_GuardMsgs_messageToStringWithoutPos___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -76,6 +80,8 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_GuardMsg static lean_object* l_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___lambda__1___closed__3; static lean_object* l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace___closed__3; +LEAN_EXPORT uint8_t l_Lean_Elab_Tactic_GuardMsgs_equalUpToNewlines(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__15; size_t lean_usize_of_nat(lean_object*); uint8_t l_String_isEmpty(lean_object*); @@ -98,6 +104,7 @@ lean_object* l_Lean_MessageData_ofSyntax(lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_GuardMsgs_parseGuardMsgsSpec___spec__6___lambda__8___closed__9; static lean_object* l_Lean_Elab_Tactic_GuardMsgs_parseGuardMsgsSpec___closed__2; static lean_object* l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__10; +static lean_object* l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace___closed__6; lean_object* l_Lean_MessageLog_toList(lean_object*); lean_object* l_Lean_Elab_Command_getRef(lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs_declRange___closed__3; @@ -110,9 +117,8 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_GuardMsg lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___lambda__1___closed__7; static lean_object* l_Lean_Elab_Tactic_GuardMsgs_parseGuardMsgsSpec___closed__4; -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare____x40_Lean_Elab_GuardMsgs___hyg_2251_(lean_object*); lean_object* l_Lean_Syntax_getOptional_x3f(lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare____x40_Lean_Elab_GuardMsgs___hyg_2251____closed__2; +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare____x40_Lean_Elab_GuardMsgs___hyg_2285_(lean_object*); static lean_object* l_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___lambda__1___closed__1; lean_object* l_EStateM_pure___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__8; @@ -126,6 +132,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs_decl LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_GuardMsgs_parseGuardMsgsSpec___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_getDocStringText___at_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_GuardMsgs_equalUpToNewlines___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__18; LEAN_EXPORT uint8_t l_Lean_Elab_Tactic_GuardMsgs_parseGuardMsgsSpec___lambda__2(lean_object*); static lean_object* l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__13; @@ -140,12 +147,14 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_GuardMsg LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Tactic_GuardMsgs_parseGuardMsgsSpec___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___spec__1___rarg(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___closed__1; +static lean_object* l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace___closed__2; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_GuardMsgs_parseGuardMsgsSpec___spec__6___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_GuardMsgs_SpecResult_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Tactic_GuardMsgs_parseGuardMsgsSpec___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_GuardMsgs_instImpl____x40_Lean_Elab_GuardMsgs___hyg_1100____closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_GuardMsgs_SpecResult_noConfusion___rarg(uint8_t, uint8_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_GuardMsgs_removeTrailingWhitespaceMarker(lean_object*); static lean_object* l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__3; LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_getDocStringText___at_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___spec__4___closed__3; @@ -172,6 +181,7 @@ uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_GuardMsgs_instImpl____x40_Lean_Elab_GuardMsgs___hyg_1100____closed__1; static lean_object* l_Lean_PersistentArray_findSomeMAux___at_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___spec__2___closed__3; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace___closed__4; uint8_t lean_uint32_dec_eq(uint32_t, uint32_t); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___closed__2; @@ -182,7 +192,6 @@ static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_GuardMsgs_par static lean_object* l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs_declRange___closed__4; uint8_t l_String_isPrefixOf(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_GuardMsgs_parseGuardMsgsSpec___lambda__1(lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare____x40_Lean_Elab_GuardMsgs___hyg_2251____closed__1; uint8_t l_Lean_Syntax_isNone(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_GuardMsgs_parseGuardMsgsSpec___elambda__1___boxed(lean_object*); static lean_object* l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__17; @@ -203,6 +212,7 @@ static lean_object* l_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___lambda__1___clo LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___spec__5(lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* l_Lean_throwError___at_Lean_Elab_Command_expandDeclId___spec__18(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_GuardMsgs_0__Lean_Elab_Tactic_GuardMsgs_messageToStringWithoutPos___lambda__1(lean_object*, lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare____x40_Lean_Elab_GuardMsgs___hyg_2285____closed__2; static lean_object* l_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___closed__2; lean_object* l_List_reverse___rarg(lean_object*); lean_object* l_String_intercalate(lean_object*, lean_object*); @@ -249,6 +259,8 @@ static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_GuardMsgs_par static lean_object* l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_GuardMsgs_instImpl____x40_Lean_Elab_GuardMsgs___hyg_1100_; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace___boxed(lean_object*); +static lean_object* l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_GuardMsgs_instTypeNameGuardMsgFailure; lean_object* l___private_Init_Dynamic_0__Dynamic_get_x3fImpl___rarg(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs_declRange___closed__5; @@ -1948,6 +1960,125 @@ x_1 = l_Lean_Elab_Tactic_GuardMsgs_instImpl____x40_Lean_Elab_GuardMsgs___hyg_110 return x_1; } } +static lean_object* _init_l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("⏎\n", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("⏎⏎\n", 7); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("\t\n", 2); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("\t⏎\n", 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(" \n", 2); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(" ⏎\n", 5); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace(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; +x_2 = l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace___closed__1; +x_3 = l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace___closed__2; +x_4 = l_String_replace(x_1, x_2, x_3); +x_5 = l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace___closed__3; +x_6 = l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace___closed__4; +x_7 = l_String_replace(x_4, x_5, x_6); +lean_dec(x_4); +x_8 = l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace___closed__5; +x_9 = l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace___closed__6; +x_10 = l_String_replace(x_7, x_8, x_9); +lean_dec(x_7); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_GuardMsgs_removeTrailingWhitespaceMarker(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace___closed__1; +x_3 = l___private_Lean_Elab_GuardMsgs_0__Lean_Elab_Tactic_GuardMsgs_messageToStringWithoutPos___lambda__2___closed__2; +x_4 = l_String_replace(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_GuardMsgs_removeTrailingWhitespaceMarker___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Elab_Tactic_GuardMsgs_removeTrailingWhitespaceMarker(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT uint8_t l_Lean_Elab_Tactic_GuardMsgs_equalUpToNewlines(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_3 = l___private_Lean_Elab_GuardMsgs_0__Lean_Elab_Tactic_GuardMsgs_messageToStringWithoutPos___lambda__2___closed__2; +x_4 = l___private_Lean_Elab_GuardMsgs_0__Lean_Elab_Tactic_GuardMsgs_messageToStringWithoutPos___lambda__4___closed__1; +x_5 = l_String_replace(x_1, x_3, x_4); +x_6 = l_String_replace(x_2, x_3, x_4); +x_7 = lean_string_dec_eq(x_5, x_6); +lean_dec(x_6); +lean_dec(x_5); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_GuardMsgs_equalUpToNewlines___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_Elab_Tactic_GuardMsgs_equalUpToNewlines(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} static lean_object* _init_l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___spec__1___rarg___closed__1() { _start: { @@ -2508,162 +2639,162 @@ return x_2; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___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; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_269; +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_264; x_7 = lean_unsigned_to_nat(1u); x_8 = l_Lean_Syntax_getArg(x_1, x_7); x_9 = lean_unsigned_to_nat(2u); x_10 = l_Lean_Syntax_getArg(x_1, x_9); x_11 = lean_unsigned_to_nat(4u); x_12 = l_Lean_Syntax_getArg(x_1, x_11); -x_269 = l_Lean_Syntax_getOptional_x3f(x_10); +x_264 = l_Lean_Syntax_getOptional_x3f(x_10); lean_dec(x_10); -if (lean_obj_tag(x_269) == 0) +if (lean_obj_tag(x_264) == 0) { -lean_object* x_270; -x_270 = lean_box(0); -x_13 = x_270; -goto block_268; +lean_object* x_265; +x_265 = lean_box(0); +x_13 = x_265; +goto block_263; } else { -uint8_t x_271; -x_271 = !lean_is_exclusive(x_269); -if (x_271 == 0) +uint8_t x_266; +x_266 = !lean_is_exclusive(x_264); +if (x_266 == 0) { -x_13 = x_269; -goto block_268; +x_13 = x_264; +goto block_263; } else { -lean_object* x_272; lean_object* x_273; -x_272 = lean_ctor_get(x_269, 0); -lean_inc(x_272); -lean_dec(x_269); -x_273 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_273, 0, x_272); -x_13 = x_273; -goto block_268; +lean_object* x_267; lean_object* x_268; +x_267 = lean_ctor_get(x_264, 0); +lean_inc(x_267); +lean_dec(x_264); +x_268 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_268, 0, x_267); +x_13 = x_268; +goto block_263; } } -block_268: +block_263: { lean_object* x_14; lean_object* x_15; if (lean_obj_tag(x_3) == 0) { -lean_object* x_249; -x_249 = lean_box(0); -x_14 = x_249; +lean_object* x_244; +x_244 = lean_box(0); +x_14 = x_244; x_15 = x_6; -goto block_248; +goto block_243; } else { -uint8_t x_250; -x_250 = !lean_is_exclusive(x_3); -if (x_250 == 0) +uint8_t x_245; +x_245 = !lean_is_exclusive(x_3); +if (x_245 == 0) { -lean_object* x_251; lean_object* x_252; -x_251 = lean_ctor_get(x_3, 0); +lean_object* x_246; lean_object* x_247; +x_246 = lean_ctor_get(x_3, 0); lean_inc(x_5); lean_inc(x_4); -x_252 = l_Lean_getDocStringText___at_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___spec__4(x_251, x_4, x_5, x_6); -if (lean_obj_tag(x_252) == 0) -{ -lean_object* x_253; lean_object* x_254; -x_253 = lean_ctor_get(x_252, 0); -lean_inc(x_253); -x_254 = lean_ctor_get(x_252, 1); -lean_inc(x_254); -lean_dec(x_252); -lean_ctor_set(x_3, 0, x_253); +x_247 = l_Lean_getDocStringText___at_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___spec__4(x_246, x_4, x_5, x_6); +if (lean_obj_tag(x_247) == 0) +{ +lean_object* x_248; lean_object* x_249; +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); +lean_ctor_set(x_3, 0, x_248); x_14 = x_3; -x_15 = x_254; -goto block_248; +x_15 = x_249; +goto block_243; } else { -uint8_t x_255; +uint8_t x_250; lean_free_object(x_3); lean_dec(x_13); lean_dec(x_12); lean_dec(x_8); lean_dec(x_5); lean_dec(x_4); -x_255 = !lean_is_exclusive(x_252); -if (x_255 == 0) +x_250 = !lean_is_exclusive(x_247); +if (x_250 == 0) { -return x_252; +return x_247; } else { -lean_object* x_256; lean_object* x_257; lean_object* x_258; -x_256 = lean_ctor_get(x_252, 0); -x_257 = lean_ctor_get(x_252, 1); -lean_inc(x_257); -lean_inc(x_256); -lean_dec(x_252); -x_258 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_258, 0, x_256); -lean_ctor_set(x_258, 1, x_257); -return x_258; +lean_object* x_251; lean_object* x_252; lean_object* x_253; +x_251 = lean_ctor_get(x_247, 0); +x_252 = lean_ctor_get(x_247, 1); +lean_inc(x_252); +lean_inc(x_251); +lean_dec(x_247); +x_253 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_253, 0, x_251); +lean_ctor_set(x_253, 1, x_252); +return x_253; } } } else { -lean_object* x_259; lean_object* x_260; -x_259 = lean_ctor_get(x_3, 0); -lean_inc(x_259); +lean_object* x_254; lean_object* x_255; +x_254 = lean_ctor_get(x_3, 0); +lean_inc(x_254); lean_dec(x_3); lean_inc(x_5); lean_inc(x_4); -x_260 = l_Lean_getDocStringText___at_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___spec__4(x_259, x_4, x_5, x_6); -if (lean_obj_tag(x_260) == 0) +x_255 = l_Lean_getDocStringText___at_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___spec__4(x_254, x_4, x_5, x_6); +if (lean_obj_tag(x_255) == 0) { -lean_object* x_261; lean_object* x_262; lean_object* x_263; -x_261 = lean_ctor_get(x_260, 0); -lean_inc(x_261); -x_262 = lean_ctor_get(x_260, 1); -lean_inc(x_262); -lean_dec(x_260); -x_263 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_263, 0, x_261); -x_14 = x_263; -x_15 = x_262; -goto block_248; +lean_object* x_256; lean_object* x_257; lean_object* x_258; +x_256 = lean_ctor_get(x_255, 0); +lean_inc(x_256); +x_257 = lean_ctor_get(x_255, 1); +lean_inc(x_257); +lean_dec(x_255); +x_258 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_258, 0, x_256); +x_14 = x_258; +x_15 = x_257; +goto block_243; } else { -lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; +lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_dec(x_13); lean_dec(x_12); lean_dec(x_8); lean_dec(x_5); lean_dec(x_4); -x_264 = lean_ctor_get(x_260, 0); -lean_inc(x_264); -x_265 = lean_ctor_get(x_260, 1); -lean_inc(x_265); -if (lean_is_exclusive(x_260)) { - lean_ctor_release(x_260, 0); - lean_ctor_release(x_260, 1); - x_266 = x_260; +x_259 = lean_ctor_get(x_255, 0); +lean_inc(x_259); +x_260 = lean_ctor_get(x_255, 1); +lean_inc(x_260); +if (lean_is_exclusive(x_255)) { + lean_ctor_release(x_255, 0); + lean_ctor_release(x_255, 1); + x_261 = x_255; } else { - lean_dec_ref(x_260); - x_266 = lean_box(0); + lean_dec_ref(x_255); + x_261 = lean_box(0); } -if (lean_is_scalar(x_266)) { - x_267 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_261)) { + x_262 = lean_alloc_ctor(1, 2, 0); } else { - x_267 = x_266; + x_262 = x_261; } -lean_ctor_set(x_267, 0, x_264); -lean_ctor_set(x_267, 1, x_265); -return x_267; +lean_ctor_set(x_262, 0, x_259); +lean_ctor_set(x_262, 1, x_260); +return x_262; } } } -block_248: +block_243: { lean_object* x_16; lean_object* x_17; lean_inc(x_5); @@ -2671,25 +2802,25 @@ lean_inc(x_4); x_16 = l_Lean_Elab_Tactic_GuardMsgs_parseGuardMsgsSpec(x_13, x_4, x_5, x_15); if (lean_obj_tag(x_14) == 0) { -lean_object* x_246; -x_246 = l___private_Lean_Elab_GuardMsgs_0__Lean_Elab_Tactic_GuardMsgs_messageToStringWithoutPos___closed__1; -x_17 = x_246; -goto block_245; +lean_object* x_241; +x_241 = l___private_Lean_Elab_GuardMsgs_0__Lean_Elab_Tactic_GuardMsgs_messageToStringWithoutPos___closed__1; +x_17 = x_241; +goto block_240; } else { -lean_object* x_247; -x_247 = lean_ctor_get(x_14, 0); -lean_inc(x_247); +lean_object* x_242; +x_242 = lean_ctor_get(x_14, 0); +lean_inc(x_242); lean_dec(x_14); -x_17 = x_247; -goto block_245; +x_17 = x_242; +goto block_240; } -block_245: +block_240: { if (lean_obj_tag(x_16) == 0) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; x_18 = lean_ctor_get(x_16, 0); lean_inc(x_18); x_19 = lean_ctor_get(x_16, 1); @@ -2697,147 +2828,144 @@ lean_inc(x_19); lean_dec(x_16); x_20 = l_String_trim(x_17); lean_dec(x_17); -x_21 = lean_st_ref_take(x_5, x_19); -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, 1); -x_26 = l_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___lambda__1___closed__3; -lean_ctor_set(x_22, 1, x_26); -x_27 = lean_st_ref_set(x_5, x_22, x_23); -x_28 = lean_ctor_get(x_27, 1); -lean_inc(x_28); -lean_dec(x_27); +x_21 = l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace___closed__1; +x_22 = l___private_Lean_Elab_GuardMsgs_0__Lean_Elab_Tactic_GuardMsgs_messageToStringWithoutPos___lambda__2___closed__2; +x_23 = l_String_replace(x_20, x_21, x_22); +lean_dec(x_20); +x_24 = lean_st_ref_take(x_5, x_19); +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_is_exclusive(x_25); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_28 = lean_ctor_get(x_25, 1); +x_29 = l_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___lambda__1___closed__3; +lean_ctor_set(x_25, 1, x_29); +x_30 = lean_st_ref_set(x_5, x_25, x_26); +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +lean_dec(x_30); lean_inc(x_5); lean_inc(x_4); -x_29 = l_Lean_Elab_Command_elabCommandTopLevel(x_12, x_4, x_5, x_28); -if (lean_obj_tag(x_29) == 0) +x_32 = l_Lean_Elab_Command_elabCommandTopLevel(x_12, x_4, x_5, x_31); +if (lean_obj_tag(x_32) == 0) { -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_30 = lean_ctor_get(x_29, 1); -lean_inc(x_30); -lean_dec(x_29); -x_31 = lean_st_ref_get(x_5, x_30); -x_32 = lean_ctor_get(x_31, 0); -lean_inc(x_32); -x_33 = lean_ctor_get(x_31, 1); +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_33 = lean_ctor_get(x_32, 1); lean_inc(x_33); -lean_dec(x_31); -x_34 = lean_ctor_get(x_32, 1); -lean_inc(x_34); lean_dec(x_32); -lean_inc(x_34); -x_35 = l_Lean_MessageLog_toList(x_34); -x_36 = l_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___lambda__1___closed__4; -x_37 = l_List_forIn_loop___at_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___spec__2(x_18, x_35, x_36, x_4, x_5, x_33); -x_38 = lean_ctor_get(x_37, 0); -lean_inc(x_38); -x_39 = lean_ctor_get(x_37, 1); -lean_inc(x_39); -lean_dec(x_37); -x_40 = lean_ctor_get(x_38, 0); -lean_inc(x_40); -x_41 = lean_ctor_get(x_38, 1); +x_34 = lean_st_ref_get(x_5, x_33); +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_34, 1); +lean_inc(x_36); +lean_dec(x_34); +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +lean_inc(x_37); +x_38 = l_Lean_MessageLog_toList(x_37); +x_39 = l_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___lambda__1___closed__4; +x_40 = l_List_forIn_loop___at_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___spec__2(x_18, x_38, x_39, x_4, x_5, x_36); +x_41 = lean_ctor_get(x_40, 0); lean_inc(x_41); -lean_dec(x_38); -x_42 = l_Lean_MessageLog_toList(x_40); -x_43 = lean_box(0); -x_44 = l_List_mapM_loop___at_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___spec__3(x_42, x_43, x_4, x_5, 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; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; -x_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_Elab_Tactic_GuardMsgs_elabGuardMsgs___lambda__1___closed__5; -x_48 = l_String_intercalate(x_47, x_45); -x_49 = l_String_trim(x_48); -lean_dec(x_48); -x_50 = l___private_Lean_Elab_GuardMsgs_0__Lean_Elab_Tactic_GuardMsgs_messageToStringWithoutPos___lambda__2___closed__2; -x_51 = l___private_Lean_Elab_GuardMsgs_0__Lean_Elab_Tactic_GuardMsgs_messageToStringWithoutPos___lambda__4___closed__1; -x_52 = l_String_replace(x_20, x_50, x_51); -lean_dec(x_20); -x_53 = l_String_replace(x_49, x_50, x_51); -x_54 = lean_string_dec_eq(x_52, x_53); -lean_dec(x_53); -lean_dec(x_52); -if (x_54 == 0) -{ -lean_object* x_55; lean_object* x_56; lean_object* x_57; uint8_t x_58; +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); +x_44 = lean_ctor_get(x_41, 1); +lean_inc(x_44); lean_dec(x_41); -x_55 = lean_st_ref_take(x_5, x_46); -x_56 = lean_ctor_get(x_55, 0); +x_45 = l_Lean_MessageLog_toList(x_43); +x_46 = lean_box(0); +x_47 = l_List_mapM_loop___at_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___spec__3(x_45, x_46, x_4, x_5, x_42); +if (lean_obj_tag(x_47) == 0) +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; uint8_t x_53; +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_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___lambda__1___closed__5; +x_51 = l_String_intercalate(x_50, x_48); +x_52 = l_String_trim(x_51); +lean_dec(x_51); +x_53 = l_Lean_Elab_Tactic_GuardMsgs_equalUpToNewlines(x_23, x_52); +lean_dec(x_23); +if (x_53 == 0) +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; uint8_t x_57; +lean_dec(x_44); +x_54 = lean_st_ref_take(x_5, x_49); +x_55 = lean_ctor_get(x_54, 0); +lean_inc(x_55); +x_56 = lean_ctor_get(x_54, 1); lean_inc(x_56); -x_57 = lean_ctor_get(x_55, 1); -lean_inc(x_57); -lean_dec(x_55); -x_58 = !lean_is_exclusive(x_56); -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; uint8_t x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; -x_59 = lean_ctor_get(x_56, 1); -lean_dec(x_59); -x_60 = l_Lean_PersistentArray_append___rarg(x_25, x_34); -lean_ctor_set(x_56, 1, x_60); -x_61 = lean_st_ref_set(x_5, x_56, x_57); -x_62 = lean_ctor_get(x_61, 1); -lean_inc(x_62); -lean_dec(x_61); -x_63 = l_Lean_stringToMessageData(x_49); -x_64 = l_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___lambda__1___closed__7; -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_getDocStringText___at_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___spec__4___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 = 2; -x_69 = l_Lean_logAt___at_Lean_Elab_Command_elabCommand___spec__4(x_8, x_67, x_68, x_4, x_5, x_62); +lean_dec(x_54); +x_57 = !lean_is_exclusive(x_55); +if (x_57 == 0) +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_58 = lean_ctor_get(x_55, 1); +lean_dec(x_58); +x_59 = l_Lean_PersistentArray_append___rarg(x_28, x_37); +lean_ctor_set(x_55, 1, x_59); +x_60 = lean_st_ref_set(x_5, x_55, x_56); +x_61 = lean_ctor_get(x_60, 1); +lean_inc(x_61); +lean_dec(x_60); +x_62 = l_Lean_stringToMessageData(x_52); +x_63 = l_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___lambda__1___closed__7; +x_64 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_64, 0, x_63); +lean_ctor_set(x_64, 1, x_62); +x_65 = l_Lean_getDocStringText___at_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___spec__4___closed__3; +x_66 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_66, 0, x_64); +lean_ctor_set(x_66, 1, x_65); +x_67 = 2; +x_68 = l_Lean_logAt___at_Lean_Elab_Command_elabCommand___spec__4(x_8, x_66, x_67, x_4, x_5, x_61); lean_dec(x_8); -x_70 = lean_ctor_get(x_69, 1); -lean_inc(x_70); -lean_dec(x_69); -x_71 = l_Lean_Elab_Command_getRef(x_4, x_5, x_70); -x_72 = lean_ctor_get(x_71, 0); +x_69 = lean_ctor_get(x_68, 1); +lean_inc(x_69); +lean_dec(x_68); +x_70 = l_Lean_Elab_Command_getRef(x_4, x_5, x_69); +x_71 = lean_ctor_get(x_70, 0); +lean_inc(x_71); +x_72 = lean_ctor_get(x_70, 1); lean_inc(x_72); -x_73 = lean_ctor_get(x_71, 1); -lean_inc(x_73); -lean_dec(x_71); -x_74 = l_Lean_Elab_Tactic_GuardMsgs_instTypeNameGuardMsgFailure; +lean_dec(x_70); +x_73 = l_Lean_Elab_Tactic_GuardMsgs_instTypeNameGuardMsgFailure; +x_74 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_74, 0, x_73); +lean_ctor_set(x_74, 1, x_52); x_75 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_75, 0, x_74); -lean_ctor_set(x_75, 1, x_49); -x_76 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_76, 0, x_72); -lean_ctor_set(x_76, 1, x_75); -x_77 = lean_alloc_ctor(8, 1, 0); -lean_ctor_set(x_77, 0, x_76); -x_78 = l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Command_expandDeclId___spec__11(x_77, x_4, x_5, x_73); +lean_ctor_set(x_75, 0, x_71); +lean_ctor_set(x_75, 1, x_74); +x_76 = lean_alloc_ctor(8, 1, 0); +lean_ctor_set(x_76, 0, x_75); +x_77 = l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Command_expandDeclId___spec__11(x_76, x_4, x_5, x_72); lean_dec(x_5); lean_dec(x_4); -return x_78; +return x_77; } 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; 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; uint8_t x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; -x_79 = lean_ctor_get(x_56, 0); -x_80 = lean_ctor_get(x_56, 2); -x_81 = lean_ctor_get(x_56, 3); -x_82 = lean_ctor_get(x_56, 4); -x_83 = lean_ctor_get(x_56, 5); -x_84 = lean_ctor_get(x_56, 6); -x_85 = lean_ctor_get(x_56, 7); -x_86 = lean_ctor_get(x_56, 8); -lean_inc(x_86); +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; uint8_t x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; +x_78 = lean_ctor_get(x_55, 0); +x_79 = lean_ctor_get(x_55, 2); +x_80 = lean_ctor_get(x_55, 3); +x_81 = lean_ctor_get(x_55, 4); +x_82 = lean_ctor_get(x_55, 5); +x_83 = lean_ctor_get(x_55, 6); +x_84 = lean_ctor_get(x_55, 7); +x_85 = lean_ctor_get(x_55, 8); lean_inc(x_85); lean_inc(x_84); lean_inc(x_83); @@ -2845,116 +2973,116 @@ lean_inc(x_82); lean_inc(x_81); lean_inc(x_80); lean_inc(x_79); -lean_dec(x_56); -x_87 = l_Lean_PersistentArray_append___rarg(x_25, x_34); -x_88 = lean_alloc_ctor(0, 9, 0); -lean_ctor_set(x_88, 0, x_79); -lean_ctor_set(x_88, 1, x_87); -lean_ctor_set(x_88, 2, x_80); -lean_ctor_set(x_88, 3, x_81); -lean_ctor_set(x_88, 4, x_82); -lean_ctor_set(x_88, 5, x_83); -lean_ctor_set(x_88, 6, x_84); -lean_ctor_set(x_88, 7, x_85); -lean_ctor_set(x_88, 8, x_86); -x_89 = lean_st_ref_set(x_5, x_88, x_57); -x_90 = lean_ctor_get(x_89, 1); -lean_inc(x_90); -lean_dec(x_89); -x_91 = l_Lean_stringToMessageData(x_49); -x_92 = l_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___lambda__1___closed__7; -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_getDocStringText___at_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___spec__4___closed__3; -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 = 2; -x_97 = l_Lean_logAt___at_Lean_Elab_Command_elabCommand___spec__4(x_8, x_95, x_96, x_4, x_5, x_90); +lean_inc(x_78); +lean_dec(x_55); +x_86 = l_Lean_PersistentArray_append___rarg(x_28, x_37); +x_87 = lean_alloc_ctor(0, 9, 0); +lean_ctor_set(x_87, 0, x_78); +lean_ctor_set(x_87, 1, x_86); +lean_ctor_set(x_87, 2, x_79); +lean_ctor_set(x_87, 3, x_80); +lean_ctor_set(x_87, 4, x_81); +lean_ctor_set(x_87, 5, x_82); +lean_ctor_set(x_87, 6, x_83); +lean_ctor_set(x_87, 7, x_84); +lean_ctor_set(x_87, 8, x_85); +x_88 = lean_st_ref_set(x_5, x_87, x_56); +x_89 = lean_ctor_get(x_88, 1); +lean_inc(x_89); +lean_dec(x_88); +x_90 = l_Lean_stringToMessageData(x_52); +x_91 = l_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___lambda__1___closed__7; +x_92 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_92, 0, x_91); +lean_ctor_set(x_92, 1, x_90); +x_93 = l_Lean_getDocStringText___at_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___spec__4___closed__3; +x_94 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_94, 0, x_92); +lean_ctor_set(x_94, 1, x_93); +x_95 = 2; +x_96 = l_Lean_logAt___at_Lean_Elab_Command_elabCommand___spec__4(x_8, x_94, x_95, x_4, x_5, x_89); lean_dec(x_8); -x_98 = lean_ctor_get(x_97, 1); -lean_inc(x_98); -lean_dec(x_97); -x_99 = l_Lean_Elab_Command_getRef(x_4, x_5, x_98); -x_100 = lean_ctor_get(x_99, 0); +x_97 = lean_ctor_get(x_96, 1); +lean_inc(x_97); +lean_dec(x_96); +x_98 = l_Lean_Elab_Command_getRef(x_4, x_5, x_97); +x_99 = lean_ctor_get(x_98, 0); +lean_inc(x_99); +x_100 = lean_ctor_get(x_98, 1); lean_inc(x_100); -x_101 = lean_ctor_get(x_99, 1); -lean_inc(x_101); -lean_dec(x_99); -x_102 = l_Lean_Elab_Tactic_GuardMsgs_instTypeNameGuardMsgFailure; +lean_dec(x_98); +x_101 = l_Lean_Elab_Tactic_GuardMsgs_instTypeNameGuardMsgFailure; +x_102 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_102, 0, x_101); +lean_ctor_set(x_102, 1, x_52); x_103 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_103, 0, x_102); -lean_ctor_set(x_103, 1, x_49); -x_104 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_104, 0, x_100); -lean_ctor_set(x_104, 1, x_103); -x_105 = lean_alloc_ctor(8, 1, 0); -lean_ctor_set(x_105, 0, x_104); -x_106 = l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Command_expandDeclId___spec__11(x_105, x_4, x_5, x_101); +lean_ctor_set(x_103, 0, x_99); +lean_ctor_set(x_103, 1, x_102); +x_104 = lean_alloc_ctor(8, 1, 0); +lean_ctor_set(x_104, 0, x_103); +x_105 = l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Command_expandDeclId___spec__11(x_104, x_4, x_5, x_100); lean_dec(x_5); lean_dec(x_4); -return x_106; +return x_105; } } else { -lean_object* x_107; lean_object* x_108; lean_object* x_109; uint8_t x_110; -lean_dec(x_49); -lean_dec(x_34); +lean_object* x_106; lean_object* x_107; lean_object* x_108; uint8_t x_109; +lean_dec(x_52); +lean_dec(x_37); lean_dec(x_8); lean_dec(x_4); -x_107 = lean_st_ref_take(x_5, x_46); -x_108 = lean_ctor_get(x_107, 0); +x_106 = lean_st_ref_take(x_5, x_49); +x_107 = lean_ctor_get(x_106, 0); +lean_inc(x_107); +x_108 = lean_ctor_get(x_106, 1); lean_inc(x_108); -x_109 = lean_ctor_get(x_107, 1); -lean_inc(x_109); -lean_dec(x_107); -x_110 = !lean_is_exclusive(x_108); -if (x_110 == 0) -{ -lean_object* x_111; lean_object* x_112; lean_object* x_113; uint8_t x_114; -x_111 = lean_ctor_get(x_108, 1); -lean_dec(x_111); -x_112 = l_Lean_PersistentArray_append___rarg(x_25, x_41); -lean_ctor_set(x_108, 1, x_112); -x_113 = lean_st_ref_set(x_5, x_108, x_109); +lean_dec(x_106); +x_109 = !lean_is_exclusive(x_107); +if (x_109 == 0) +{ +lean_object* x_110; lean_object* x_111; lean_object* x_112; uint8_t x_113; +x_110 = lean_ctor_get(x_107, 1); +lean_dec(x_110); +x_111 = l_Lean_PersistentArray_append___rarg(x_28, x_44); +lean_ctor_set(x_107, 1, x_111); +x_112 = lean_st_ref_set(x_5, x_107, x_108); lean_dec(x_5); -x_114 = !lean_is_exclusive(x_113); -if (x_114 == 0) +x_113 = !lean_is_exclusive(x_112); +if (x_113 == 0) { -lean_object* x_115; lean_object* x_116; -x_115 = lean_ctor_get(x_113, 0); -lean_dec(x_115); -x_116 = lean_box(0); -lean_ctor_set(x_113, 0, x_116); -return x_113; +lean_object* x_114; lean_object* x_115; +x_114 = lean_ctor_get(x_112, 0); +lean_dec(x_114); +x_115 = lean_box(0); +lean_ctor_set(x_112, 0, x_115); +return x_112; } else { -lean_object* x_117; lean_object* x_118; lean_object* x_119; -x_117 = lean_ctor_get(x_113, 1); -lean_inc(x_117); -lean_dec(x_113); -x_118 = lean_box(0); -x_119 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_119, 0, x_118); -lean_ctor_set(x_119, 1, x_117); -return x_119; +lean_object* x_116; lean_object* x_117; lean_object* x_118; +x_116 = lean_ctor_get(x_112, 1); +lean_inc(x_116); +lean_dec(x_112); +x_117 = lean_box(0); +x_118 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_118, 0, x_117); +lean_ctor_set(x_118, 1, x_116); +return x_118; } } else { -lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; -x_120 = lean_ctor_get(x_108, 0); -x_121 = lean_ctor_get(x_108, 2); -x_122 = lean_ctor_get(x_108, 3); -x_123 = lean_ctor_get(x_108, 4); -x_124 = lean_ctor_get(x_108, 5); -x_125 = lean_ctor_get(x_108, 6); -x_126 = lean_ctor_get(x_108, 7); -x_127 = lean_ctor_get(x_108, 8); -lean_inc(x_127); +lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; +x_119 = lean_ctor_get(x_107, 0); +x_120 = lean_ctor_get(x_107, 2); +x_121 = lean_ctor_get(x_107, 3); +x_122 = lean_ctor_get(x_107, 4); +x_123 = lean_ctor_get(x_107, 5); +x_124 = lean_ctor_get(x_107, 6); +x_125 = lean_ctor_get(x_107, 7); +x_126 = lean_ctor_get(x_107, 8); lean_inc(x_126); lean_inc(x_125); lean_inc(x_124); @@ -2962,114 +3090,114 @@ lean_inc(x_123); lean_inc(x_122); lean_inc(x_121); lean_inc(x_120); -lean_dec(x_108); -x_128 = l_Lean_PersistentArray_append___rarg(x_25, x_41); -x_129 = lean_alloc_ctor(0, 9, 0); -lean_ctor_set(x_129, 0, x_120); -lean_ctor_set(x_129, 1, x_128); -lean_ctor_set(x_129, 2, x_121); -lean_ctor_set(x_129, 3, x_122); -lean_ctor_set(x_129, 4, x_123); -lean_ctor_set(x_129, 5, x_124); -lean_ctor_set(x_129, 6, x_125); -lean_ctor_set(x_129, 7, x_126); -lean_ctor_set(x_129, 8, x_127); -x_130 = lean_st_ref_set(x_5, x_129, x_109); +lean_inc(x_119); +lean_dec(x_107); +x_127 = l_Lean_PersistentArray_append___rarg(x_28, x_44); +x_128 = lean_alloc_ctor(0, 9, 0); +lean_ctor_set(x_128, 0, x_119); +lean_ctor_set(x_128, 1, x_127); +lean_ctor_set(x_128, 2, x_120); +lean_ctor_set(x_128, 3, x_121); +lean_ctor_set(x_128, 4, x_122); +lean_ctor_set(x_128, 5, x_123); +lean_ctor_set(x_128, 6, x_124); +lean_ctor_set(x_128, 7, x_125); +lean_ctor_set(x_128, 8, x_126); +x_129 = lean_st_ref_set(x_5, x_128, x_108); lean_dec(x_5); -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; +x_130 = lean_ctor_get(x_129, 1); +lean_inc(x_130); +if (lean_is_exclusive(x_129)) { + lean_ctor_release(x_129, 0); + lean_ctor_release(x_129, 1); + x_131 = x_129; } else { - lean_dec_ref(x_130); - x_132 = lean_box(0); + lean_dec_ref(x_129); + x_131 = lean_box(0); } -x_133 = lean_box(0); -if (lean_is_scalar(x_132)) { - x_134 = lean_alloc_ctor(0, 2, 0); +x_132 = lean_box(0); +if (lean_is_scalar(x_131)) { + x_133 = lean_alloc_ctor(0, 2, 0); } else { - x_134 = x_132; + x_133 = x_131; } -lean_ctor_set(x_134, 0, x_133); -lean_ctor_set(x_134, 1, x_131); -return x_134; +lean_ctor_set(x_133, 0, x_132); +lean_ctor_set(x_133, 1, x_130); +return x_133; } } } else { -uint8_t x_135; -lean_dec(x_41); -lean_dec(x_34); -lean_dec(x_25); -lean_dec(x_20); +uint8_t x_134; +lean_dec(x_44); +lean_dec(x_37); +lean_dec(x_28); +lean_dec(x_23); lean_dec(x_8); lean_dec(x_5); lean_dec(x_4); -x_135 = !lean_is_exclusive(x_44); -if (x_135 == 0) +x_134 = !lean_is_exclusive(x_47); +if (x_134 == 0) { -return x_44; +return x_47; } else { -lean_object* x_136; lean_object* x_137; lean_object* x_138; -x_136 = lean_ctor_get(x_44, 0); -x_137 = lean_ctor_get(x_44, 1); -lean_inc(x_137); +lean_object* x_135; lean_object* x_136; lean_object* x_137; +x_135 = lean_ctor_get(x_47, 0); +x_136 = lean_ctor_get(x_47, 1); lean_inc(x_136); -lean_dec(x_44); -x_138 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_138, 0, x_136); -lean_ctor_set(x_138, 1, x_137); -return x_138; +lean_inc(x_135); +lean_dec(x_47); +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; } } } else { -uint8_t x_139; -lean_dec(x_25); -lean_dec(x_20); +uint8_t x_138; +lean_dec(x_28); +lean_dec(x_23); lean_dec(x_18); lean_dec(x_8); lean_dec(x_5); lean_dec(x_4); -x_139 = !lean_is_exclusive(x_29); -if (x_139 == 0) +x_138 = !lean_is_exclusive(x_32); +if (x_138 == 0) { -return x_29; +return x_32; } else { -lean_object* x_140; lean_object* x_141; lean_object* x_142; -x_140 = lean_ctor_get(x_29, 0); -x_141 = lean_ctor_get(x_29, 1); -lean_inc(x_141); +lean_object* x_139; lean_object* x_140; lean_object* x_141; +x_139 = lean_ctor_get(x_32, 0); +x_140 = lean_ctor_get(x_32, 1); lean_inc(x_140); -lean_dec(x_29); -x_142 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_142, 0, x_140); -lean_ctor_set(x_142, 1, x_141); -return x_142; +lean_inc(x_139); +lean_dec(x_32); +x_141 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_141, 0, x_139); +lean_ctor_set(x_141, 1, x_140); +return x_141; } } } else { -lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; -x_143 = lean_ctor_get(x_22, 0); -x_144 = lean_ctor_get(x_22, 1); -x_145 = lean_ctor_get(x_22, 2); -x_146 = lean_ctor_get(x_22, 3); -x_147 = lean_ctor_get(x_22, 4); -x_148 = lean_ctor_get(x_22, 5); -x_149 = lean_ctor_get(x_22, 6); -x_150 = lean_ctor_get(x_22, 7); -x_151 = lean_ctor_get(x_22, 8); -lean_inc(x_151); +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_142 = lean_ctor_get(x_25, 0); +x_143 = lean_ctor_get(x_25, 1); +x_144 = lean_ctor_get(x_25, 2); +x_145 = lean_ctor_get(x_25, 3); +x_146 = lean_ctor_get(x_25, 4); +x_147 = lean_ctor_get(x_25, 5); +x_148 = lean_ctor_get(x_25, 6); +x_149 = lean_ctor_get(x_25, 7); +x_150 = lean_ctor_get(x_25, 8); lean_inc(x_150); lean_inc(x_149); lean_inc(x_148); @@ -3078,343 +3206,338 @@ lean_inc(x_146); lean_inc(x_145); lean_inc(x_144); lean_inc(x_143); -lean_dec(x_22); -x_152 = l_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___lambda__1___closed__3; -x_153 = lean_alloc_ctor(0, 9, 0); -lean_ctor_set(x_153, 0, x_143); -lean_ctor_set(x_153, 1, x_152); -lean_ctor_set(x_153, 2, x_145); -lean_ctor_set(x_153, 3, x_146); -lean_ctor_set(x_153, 4, x_147); -lean_ctor_set(x_153, 5, x_148); -lean_ctor_set(x_153, 6, x_149); -lean_ctor_set(x_153, 7, x_150); -lean_ctor_set(x_153, 8, x_151); -x_154 = lean_st_ref_set(x_5, x_153, x_23); -x_155 = lean_ctor_get(x_154, 1); -lean_inc(x_155); -lean_dec(x_154); +lean_inc(x_142); +lean_dec(x_25); +x_151 = l_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___lambda__1___closed__3; +x_152 = lean_alloc_ctor(0, 9, 0); +lean_ctor_set(x_152, 0, x_142); +lean_ctor_set(x_152, 1, x_151); +lean_ctor_set(x_152, 2, x_144); +lean_ctor_set(x_152, 3, x_145); +lean_ctor_set(x_152, 4, x_146); +lean_ctor_set(x_152, 5, x_147); +lean_ctor_set(x_152, 6, x_148); +lean_ctor_set(x_152, 7, x_149); +lean_ctor_set(x_152, 8, x_150); +x_153 = lean_st_ref_set(x_5, x_152, x_26); +x_154 = lean_ctor_get(x_153, 1); +lean_inc(x_154); +lean_dec(x_153); lean_inc(x_5); lean_inc(x_4); -x_156 = l_Lean_Elab_Command_elabCommandTopLevel(x_12, x_4, x_5, x_155); -if (lean_obj_tag(x_156) == 0) -{ -lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; 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_157 = lean_ctor_get(x_156, 1); -lean_inc(x_157); -lean_dec(x_156); -x_158 = lean_st_ref_get(x_5, x_157); -x_159 = lean_ctor_get(x_158, 0); +x_155 = l_Lean_Elab_Command_elabCommandTopLevel(x_12, x_4, x_5, x_154); +if (lean_obj_tag(x_155) == 0) +{ +lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; +x_156 = lean_ctor_get(x_155, 1); +lean_inc(x_156); +lean_dec(x_155); +x_157 = lean_st_ref_get(x_5, x_156); +x_158 = lean_ctor_get(x_157, 0); +lean_inc(x_158); +x_159 = lean_ctor_get(x_157, 1); lean_inc(x_159); +lean_dec(x_157); x_160 = lean_ctor_get(x_158, 1); lean_inc(x_160); lean_dec(x_158); -x_161 = lean_ctor_get(x_159, 1); -lean_inc(x_161); -lean_dec(x_159); -lean_inc(x_161); -x_162 = l_Lean_MessageLog_toList(x_161); -x_163 = l_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___lambda__1___closed__4; -x_164 = l_List_forIn_loop___at_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___spec__2(x_18, x_162, x_163, x_4, x_5, x_160); -x_165 = lean_ctor_get(x_164, 0); +lean_inc(x_160); +x_161 = l_Lean_MessageLog_toList(x_160); +x_162 = l_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___lambda__1___closed__4; +x_163 = l_List_forIn_loop___at_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___spec__2(x_18, x_161, x_162, x_4, x_5, x_159); +x_164 = lean_ctor_get(x_163, 0); +lean_inc(x_164); +x_165 = lean_ctor_get(x_163, 1); lean_inc(x_165); -x_166 = lean_ctor_get(x_164, 1); +lean_dec(x_163); +x_166 = lean_ctor_get(x_164, 0); lean_inc(x_166); -lean_dec(x_164); -x_167 = lean_ctor_get(x_165, 0); +x_167 = lean_ctor_get(x_164, 1); lean_inc(x_167); -x_168 = lean_ctor_get(x_165, 1); -lean_inc(x_168); -lean_dec(x_165); -x_169 = l_Lean_MessageLog_toList(x_167); -x_170 = lean_box(0); -x_171 = l_List_mapM_loop___at_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___spec__3(x_169, x_170, x_4, x_5, x_166); -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_178; lean_object* x_179; lean_object* x_180; uint8_t x_181; -x_172 = lean_ctor_get(x_171, 0); +lean_dec(x_164); +x_168 = l_Lean_MessageLog_toList(x_166); +x_169 = lean_box(0); +x_170 = l_List_mapM_loop___at_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___spec__3(x_168, x_169, x_4, x_5, x_165); +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; uint8_t 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); -x_173 = lean_ctor_get(x_171, 1); -lean_inc(x_173); -lean_dec(x_171); -x_174 = l_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___lambda__1___closed__5; -x_175 = l_String_intercalate(x_174, x_172); -x_176 = l_String_trim(x_175); -lean_dec(x_175); -x_177 = l___private_Lean_Elab_GuardMsgs_0__Lean_Elab_Tactic_GuardMsgs_messageToStringWithoutPos___lambda__2___closed__2; -x_178 = l___private_Lean_Elab_GuardMsgs_0__Lean_Elab_Tactic_GuardMsgs_messageToStringWithoutPos___lambda__4___closed__1; -x_179 = l_String_replace(x_20, x_177, x_178); -lean_dec(x_20); -x_180 = l_String_replace(x_176, x_177, x_178); -x_181 = lean_string_dec_eq(x_179, x_180); -lean_dec(x_180); -lean_dec(x_179); -if (x_181 == 0) -{ -lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; uint8_t 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_dec(x_168); -x_182 = lean_st_ref_take(x_5, x_173); -x_183 = lean_ctor_get(x_182, 0); +lean_dec(x_170); +x_173 = l_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___lambda__1___closed__5; +x_174 = l_String_intercalate(x_173, x_171); +x_175 = l_String_trim(x_174); +lean_dec(x_174); +x_176 = l_Lean_Elab_Tactic_GuardMsgs_equalUpToNewlines(x_23, x_175); +lean_dec(x_23); +if (x_176 == 0) +{ +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; uint8_t 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_dec(x_167); +x_177 = lean_st_ref_take(x_5, x_172); +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_ctor_get(x_178, 0); +lean_inc(x_180); +x_181 = lean_ctor_get(x_178, 2); +lean_inc(x_181); +x_182 = lean_ctor_get(x_178, 3); +lean_inc(x_182); +x_183 = lean_ctor_get(x_178, 4); lean_inc(x_183); -x_184 = lean_ctor_get(x_182, 1); +x_184 = lean_ctor_get(x_178, 5); lean_inc(x_184); -lean_dec(x_182); -x_185 = lean_ctor_get(x_183, 0); +x_185 = lean_ctor_get(x_178, 6); lean_inc(x_185); -x_186 = lean_ctor_get(x_183, 2); +x_186 = lean_ctor_get(x_178, 7); lean_inc(x_186); -x_187 = lean_ctor_get(x_183, 3); +x_187 = lean_ctor_get(x_178, 8); lean_inc(x_187); -x_188 = lean_ctor_get(x_183, 4); -lean_inc(x_188); -x_189 = lean_ctor_get(x_183, 5); -lean_inc(x_189); -x_190 = lean_ctor_get(x_183, 6); -lean_inc(x_190); -x_191 = lean_ctor_get(x_183, 7); -lean_inc(x_191); -x_192 = lean_ctor_get(x_183, 8); -lean_inc(x_192); -if (lean_is_exclusive(x_183)) { - lean_ctor_release(x_183, 0); - lean_ctor_release(x_183, 1); - lean_ctor_release(x_183, 2); - lean_ctor_release(x_183, 3); - lean_ctor_release(x_183, 4); - lean_ctor_release(x_183, 5); - lean_ctor_release(x_183, 6); - lean_ctor_release(x_183, 7); - lean_ctor_release(x_183, 8); - x_193 = x_183; +if (lean_is_exclusive(x_178)) { + lean_ctor_release(x_178, 0); + lean_ctor_release(x_178, 1); + lean_ctor_release(x_178, 2); + lean_ctor_release(x_178, 3); + lean_ctor_release(x_178, 4); + lean_ctor_release(x_178, 5); + lean_ctor_release(x_178, 6); + lean_ctor_release(x_178, 7); + lean_ctor_release(x_178, 8); + x_188 = x_178; } else { - lean_dec_ref(x_183); - x_193 = lean_box(0); + lean_dec_ref(x_178); + x_188 = lean_box(0); } -x_194 = l_Lean_PersistentArray_append___rarg(x_144, x_161); -if (lean_is_scalar(x_193)) { - x_195 = lean_alloc_ctor(0, 9, 0); +x_189 = l_Lean_PersistentArray_append___rarg(x_143, x_160); +if (lean_is_scalar(x_188)) { + x_190 = lean_alloc_ctor(0, 9, 0); } else { - x_195 = x_193; -} -lean_ctor_set(x_195, 0, x_185); -lean_ctor_set(x_195, 1, x_194); -lean_ctor_set(x_195, 2, x_186); -lean_ctor_set(x_195, 3, x_187); -lean_ctor_set(x_195, 4, x_188); -lean_ctor_set(x_195, 5, x_189); -lean_ctor_set(x_195, 6, x_190); -lean_ctor_set(x_195, 7, x_191); -lean_ctor_set(x_195, 8, x_192); -x_196 = lean_st_ref_set(x_5, x_195, x_184); -x_197 = lean_ctor_get(x_196, 1); -lean_inc(x_197); -lean_dec(x_196); -x_198 = l_Lean_stringToMessageData(x_176); -x_199 = l_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___lambda__1___closed__7; -x_200 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_200, 0, x_199); -lean_ctor_set(x_200, 1, x_198); -x_201 = l_Lean_getDocStringText___at_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___spec__4___closed__3; -x_202 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_202, 0, x_200); -lean_ctor_set(x_202, 1, x_201); -x_203 = 2; -x_204 = l_Lean_logAt___at_Lean_Elab_Command_elabCommand___spec__4(x_8, x_202, x_203, x_4, x_5, x_197); + x_190 = x_188; +} +lean_ctor_set(x_190, 0, x_180); +lean_ctor_set(x_190, 1, x_189); +lean_ctor_set(x_190, 2, x_181); +lean_ctor_set(x_190, 3, x_182); +lean_ctor_set(x_190, 4, x_183); +lean_ctor_set(x_190, 5, x_184); +lean_ctor_set(x_190, 6, x_185); +lean_ctor_set(x_190, 7, x_186); +lean_ctor_set(x_190, 8, x_187); +x_191 = lean_st_ref_set(x_5, x_190, x_179); +x_192 = lean_ctor_get(x_191, 1); +lean_inc(x_192); +lean_dec(x_191); +x_193 = l_Lean_stringToMessageData(x_175); +x_194 = l_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___lambda__1___closed__7; +x_195 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_195, 0, x_194); +lean_ctor_set(x_195, 1, x_193); +x_196 = l_Lean_getDocStringText___at_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___spec__4___closed__3; +x_197 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_197, 0, x_195); +lean_ctor_set(x_197, 1, x_196); +x_198 = 2; +x_199 = l_Lean_logAt___at_Lean_Elab_Command_elabCommand___spec__4(x_8, x_197, x_198, x_4, x_5, x_192); lean_dec(x_8); -x_205 = lean_ctor_get(x_204, 1); -lean_inc(x_205); -lean_dec(x_204); -x_206 = l_Lean_Elab_Command_getRef(x_4, x_5, 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_dec(x_206); -x_209 = l_Lean_Elab_Tactic_GuardMsgs_instTypeNameGuardMsgFailure; -x_210 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_210, 0, x_209); -lean_ctor_set(x_210, 1, x_176); -x_211 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_211, 0, x_207); -lean_ctor_set(x_211, 1, x_210); -x_212 = lean_alloc_ctor(8, 1, 0); -lean_ctor_set(x_212, 0, x_211); -x_213 = l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Command_expandDeclId___spec__11(x_212, x_4, x_5, x_208); +x_200 = lean_ctor_get(x_199, 1); +lean_inc(x_200); +lean_dec(x_199); +x_201 = l_Lean_Elab_Command_getRef(x_4, x_5, x_200); +x_202 = lean_ctor_get(x_201, 0); +lean_inc(x_202); +x_203 = lean_ctor_get(x_201, 1); +lean_inc(x_203); +lean_dec(x_201); +x_204 = l_Lean_Elab_Tactic_GuardMsgs_instTypeNameGuardMsgFailure; +x_205 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_205, 0, x_204); +lean_ctor_set(x_205, 1, x_175); +x_206 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_206, 0, x_202); +lean_ctor_set(x_206, 1, x_205); +x_207 = lean_alloc_ctor(8, 1, 0); +lean_ctor_set(x_207, 0, x_206); +x_208 = l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Command_expandDeclId___spec__11(x_207, x_4, x_5, x_203); lean_dec(x_5); lean_dec(x_4); -return x_213; +return x_208; } else { -lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; -lean_dec(x_176); -lean_dec(x_161); +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_dec(x_175); +lean_dec(x_160); lean_dec(x_8); lean_dec(x_4); -x_214 = lean_st_ref_take(x_5, x_173); -x_215 = lean_ctor_get(x_214, 0); +x_209 = lean_st_ref_take(x_5, x_172); +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, 2); +lean_inc(x_213); +x_214 = lean_ctor_get(x_210, 3); +lean_inc(x_214); +x_215 = lean_ctor_get(x_210, 4); lean_inc(x_215); -x_216 = lean_ctor_get(x_214, 1); +x_216 = lean_ctor_get(x_210, 5); lean_inc(x_216); -lean_dec(x_214); -x_217 = lean_ctor_get(x_215, 0); +x_217 = lean_ctor_get(x_210, 6); lean_inc(x_217); -x_218 = lean_ctor_get(x_215, 2); +x_218 = lean_ctor_get(x_210, 7); lean_inc(x_218); -x_219 = lean_ctor_get(x_215, 3); +x_219 = lean_ctor_get(x_210, 8); lean_inc(x_219); -x_220 = lean_ctor_get(x_215, 4); -lean_inc(x_220); -x_221 = lean_ctor_get(x_215, 5); -lean_inc(x_221); -x_222 = lean_ctor_get(x_215, 6); -lean_inc(x_222); -x_223 = lean_ctor_get(x_215, 7); -lean_inc(x_223); -x_224 = lean_ctor_get(x_215, 8); -lean_inc(x_224); -if (lean_is_exclusive(x_215)) { - lean_ctor_release(x_215, 0); - lean_ctor_release(x_215, 1); - lean_ctor_release(x_215, 2); - lean_ctor_release(x_215, 3); - lean_ctor_release(x_215, 4); - lean_ctor_release(x_215, 5); - lean_ctor_release(x_215, 6); - lean_ctor_release(x_215, 7); - lean_ctor_release(x_215, 8); - x_225 = x_215; +if (lean_is_exclusive(x_210)) { + lean_ctor_release(x_210, 0); + lean_ctor_release(x_210, 1); + lean_ctor_release(x_210, 2); + lean_ctor_release(x_210, 3); + lean_ctor_release(x_210, 4); + lean_ctor_release(x_210, 5); + lean_ctor_release(x_210, 6); + lean_ctor_release(x_210, 7); + lean_ctor_release(x_210, 8); + x_220 = x_210; } else { - lean_dec_ref(x_215); - x_225 = lean_box(0); + lean_dec_ref(x_210); + x_220 = lean_box(0); } -x_226 = l_Lean_PersistentArray_append___rarg(x_144, x_168); -if (lean_is_scalar(x_225)) { - x_227 = lean_alloc_ctor(0, 9, 0); +x_221 = l_Lean_PersistentArray_append___rarg(x_143, x_167); +if (lean_is_scalar(x_220)) { + x_222 = lean_alloc_ctor(0, 9, 0); } else { - x_227 = x_225; -} -lean_ctor_set(x_227, 0, x_217); -lean_ctor_set(x_227, 1, x_226); -lean_ctor_set(x_227, 2, x_218); -lean_ctor_set(x_227, 3, x_219); -lean_ctor_set(x_227, 4, x_220); -lean_ctor_set(x_227, 5, x_221); -lean_ctor_set(x_227, 6, x_222); -lean_ctor_set(x_227, 7, x_223); -lean_ctor_set(x_227, 8, x_224); -x_228 = lean_st_ref_set(x_5, x_227, x_216); + x_222 = x_220; +} +lean_ctor_set(x_222, 0, x_212); +lean_ctor_set(x_222, 1, x_221); +lean_ctor_set(x_222, 2, x_213); +lean_ctor_set(x_222, 3, x_214); +lean_ctor_set(x_222, 4, x_215); +lean_ctor_set(x_222, 5, x_216); +lean_ctor_set(x_222, 6, x_217); +lean_ctor_set(x_222, 7, x_218); +lean_ctor_set(x_222, 8, x_219); +x_223 = lean_st_ref_set(x_5, x_222, x_211); lean_dec(x_5); -x_229 = lean_ctor_get(x_228, 1); -lean_inc(x_229); -if (lean_is_exclusive(x_228)) { - lean_ctor_release(x_228, 0); - lean_ctor_release(x_228, 1); - x_230 = x_228; +x_224 = lean_ctor_get(x_223, 1); +lean_inc(x_224); +if (lean_is_exclusive(x_223)) { + lean_ctor_release(x_223, 0); + lean_ctor_release(x_223, 1); + x_225 = x_223; } else { - lean_dec_ref(x_228); - x_230 = lean_box(0); + lean_dec_ref(x_223); + x_225 = lean_box(0); } -x_231 = lean_box(0); -if (lean_is_scalar(x_230)) { - x_232 = lean_alloc_ctor(0, 2, 0); +x_226 = lean_box(0); +if (lean_is_scalar(x_225)) { + x_227 = lean_alloc_ctor(0, 2, 0); } else { - x_232 = x_230; + x_227 = x_225; } -lean_ctor_set(x_232, 0, x_231); -lean_ctor_set(x_232, 1, x_229); -return x_232; +lean_ctor_set(x_227, 0, x_226); +lean_ctor_set(x_227, 1, x_224); +return x_227; } } else { -lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; -lean_dec(x_168); -lean_dec(x_161); -lean_dec(x_144); -lean_dec(x_20); +lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; +lean_dec(x_167); +lean_dec(x_160); +lean_dec(x_143); +lean_dec(x_23); lean_dec(x_8); lean_dec(x_5); lean_dec(x_4); -x_233 = lean_ctor_get(x_171, 0); -lean_inc(x_233); -x_234 = lean_ctor_get(x_171, 1); -lean_inc(x_234); -if (lean_is_exclusive(x_171)) { - lean_ctor_release(x_171, 0); - lean_ctor_release(x_171, 1); - x_235 = x_171; +x_228 = lean_ctor_get(x_170, 0); +lean_inc(x_228); +x_229 = lean_ctor_get(x_170, 1); +lean_inc(x_229); +if (lean_is_exclusive(x_170)) { + lean_ctor_release(x_170, 0); + lean_ctor_release(x_170, 1); + x_230 = x_170; } else { - lean_dec_ref(x_171); - x_235 = lean_box(0); + lean_dec_ref(x_170); + x_230 = lean_box(0); } -if (lean_is_scalar(x_235)) { - x_236 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_230)) { + x_231 = lean_alloc_ctor(1, 2, 0); } else { - x_236 = x_235; + x_231 = x_230; } -lean_ctor_set(x_236, 0, x_233); -lean_ctor_set(x_236, 1, x_234); -return x_236; +lean_ctor_set(x_231, 0, x_228); +lean_ctor_set(x_231, 1, x_229); +return x_231; } } else { -lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; -lean_dec(x_144); -lean_dec(x_20); +lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; +lean_dec(x_143); +lean_dec(x_23); lean_dec(x_18); lean_dec(x_8); lean_dec(x_5); lean_dec(x_4); -x_237 = lean_ctor_get(x_156, 0); -lean_inc(x_237); -x_238 = lean_ctor_get(x_156, 1); -lean_inc(x_238); -if (lean_is_exclusive(x_156)) { - lean_ctor_release(x_156, 0); - lean_ctor_release(x_156, 1); - x_239 = x_156; +x_232 = lean_ctor_get(x_155, 0); +lean_inc(x_232); +x_233 = lean_ctor_get(x_155, 1); +lean_inc(x_233); +if (lean_is_exclusive(x_155)) { + lean_ctor_release(x_155, 0); + lean_ctor_release(x_155, 1); + x_234 = x_155; } else { - lean_dec_ref(x_156); - x_239 = lean_box(0); + lean_dec_ref(x_155); + x_234 = lean_box(0); } -if (lean_is_scalar(x_239)) { - x_240 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_234)) { + x_235 = lean_alloc_ctor(1, 2, 0); } else { - x_240 = x_239; + x_235 = x_234; } -lean_ctor_set(x_240, 0, x_237); -lean_ctor_set(x_240, 1, x_238); -return x_240; +lean_ctor_set(x_235, 0, x_232); +lean_ctor_set(x_235, 1, x_233); +return x_235; } } } else { -uint8_t x_241; +uint8_t x_236; lean_dec(x_17); lean_dec(x_12); lean_dec(x_8); lean_dec(x_5); lean_dec(x_4); -x_241 = !lean_is_exclusive(x_16); -if (x_241 == 0) +x_236 = !lean_is_exclusive(x_16); +if (x_236 == 0) { return x_16; } else { -lean_object* x_242; lean_object* x_243; lean_object* x_244; -x_242 = lean_ctor_get(x_16, 0); -x_243 = lean_ctor_get(x_16, 1); -lean_inc(x_243); -lean_inc(x_242); +lean_object* x_237; lean_object* x_238; lean_object* x_239; +x_237 = lean_ctor_get(x_16, 0); +x_238 = lean_ctor_get(x_16, 1); +lean_inc(x_238); +lean_inc(x_237); lean_dec(x_16); -x_244 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_244, 0, x_242); -lean_ctor_set(x_244, 1, x_243); -return x_244; +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; } } } @@ -3659,7 +3782,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_elabGuardMsg _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(76u); +x_1 = lean_unsigned_to_nat(94u); x_2 = lean_unsigned_to_nat(42u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -3671,7 +3794,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_elabGuardMsg _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(100u); +x_1 = lean_unsigned_to_nat(118u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -3699,7 +3822,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_elabGuardMsg _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(76u); +x_1 = lean_unsigned_to_nat(94u); x_2 = lean_unsigned_to_nat(46u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -3711,7 +3834,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_elabGuardMsg _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(76u); +x_1 = lean_unsigned_to_nat(94u); x_2 = lean_unsigned_to_nat(59u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -4448,397 +4571,401 @@ uint8_t x_25; x_25 = !lean_is_exclusive(x_23); if (x_25 == 0) { -lean_object* x_26; uint8_t x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +lean_object* x_26; lean_object* x_27; uint8_t 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_23, 0); -x_27 = l_String_isEmpty(x_10); -x_28 = l_Lean_Server_FileWorker_EditableDocument_versionedIdentifier(x_13); -x_29 = lean_ctor_get(x_13, 0); -lean_inc(x_29); -lean_dec(x_13); -x_30 = lean_ctor_get(x_29, 2); +x_27 = l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace(x_10); +lean_dec(x_10); +x_28 = l_String_isEmpty(x_27); +x_29 = l_Lean_Server_FileWorker_EditableDocument_versionedIdentifier(x_13); +x_30 = lean_ctor_get(x_13, 0); lean_inc(x_30); -lean_dec(x_29); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_19); -lean_ctor_set(x_31, 1, x_26); -x_32 = l_Lean_FileMap_utf8RangeToLspRange(x_30, x_31); -lean_dec(x_31); +lean_dec(x_13); +x_31 = lean_ctor_get(x_30, 2); +lean_inc(x_31); lean_dec(x_30); -if (x_27 == 0) +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_19); +lean_ctor_set(x_32, 1, x_26); +x_33 = l_Lean_FileMap_utf8RangeToLspRange(x_31, x_32); +lean_dec(x_32); +lean_dec(x_31); +if (x_28 == 0) { -lean_object* x_33; lean_object* x_34; uint8_t x_35; -x_33 = lean_string_length(x_10); -x_34 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__14; -x_35 = lean_nat_dec_le(x_33, x_34); -lean_dec(x_33); -if (x_35 == 0) +lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_34 = lean_string_length(x_27); +x_35 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__14; +x_36 = lean_nat_dec_le(x_34, x_35); +lean_dec(x_34); +if (x_36 == 0) { -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; 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 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__15; -x_37 = lean_string_append(x_36, x_10); -lean_dec(x_10); -x_38 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__16; -x_39 = lean_string_append(x_37, x_38); -x_40 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_40, 0, x_32); -lean_ctor_set(x_40, 1, x_39); -lean_ctor_set(x_40, 2, x_14); -lean_ctor_set(x_40, 3, x_14); -x_41 = l_Lean_Lsp_WorkspaceEdit_ofTextEdit(x_28, x_40); -lean_ctor_set(x_23, 0, x_41); -x_42 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__4; -x_43 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__2; -x_44 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__3; -x_45 = lean_alloc_ctor(0, 10, 0); -lean_ctor_set(x_45, 0, x_14); -lean_ctor_set(x_45, 1, x_14); -lean_ctor_set(x_45, 2, x_42); -lean_ctor_set(x_45, 3, x_43); -lean_ctor_set(x_45, 4, x_14); -lean_ctor_set(x_45, 5, x_44); -lean_ctor_set(x_45, 6, x_14); -lean_ctor_set(x_45, 7, x_23); -lean_ctor_set(x_45, 8, x_14); -lean_ctor_set(x_45, 9, x_14); -x_46 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 1); -lean_closure_set(x_46, 0, x_45); -lean_ctor_set(x_16, 0, x_46); -x_47 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__5; -x_48 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_48, 0, x_47); -lean_ctor_set(x_48, 1, x_16); -x_49 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__6; -x_50 = lean_array_push(x_49, x_48); -lean_ctor_set(x_11, 0, x_50); +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; 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_37 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__15; +x_38 = lean_string_append(x_37, x_27); +lean_dec(x_27); +x_39 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__16; +x_40 = lean_string_append(x_38, x_39); +x_41 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_41, 0, x_33); +lean_ctor_set(x_41, 1, x_40); +lean_ctor_set(x_41, 2, x_14); +lean_ctor_set(x_41, 3, x_14); +x_42 = l_Lean_Lsp_WorkspaceEdit_ofTextEdit(x_29, x_41); +lean_ctor_set(x_23, 0, x_42); +x_43 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__4; +x_44 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__2; +x_45 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__3; +x_46 = lean_alloc_ctor(0, 10, 0); +lean_ctor_set(x_46, 0, x_14); +lean_ctor_set(x_46, 1, x_14); +lean_ctor_set(x_46, 2, x_43); +lean_ctor_set(x_46, 3, x_44); +lean_ctor_set(x_46, 4, x_14); +lean_ctor_set(x_46, 5, x_45); +lean_ctor_set(x_46, 6, x_14); +lean_ctor_set(x_46, 7, x_23); +lean_ctor_set(x_46, 8, x_14); +lean_ctor_set(x_46, 9, x_14); +x_47 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 1); +lean_closure_set(x_47, 0, x_46); +lean_ctor_set(x_16, 0, x_47); +x_48 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__5; +x_49 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_16); +x_50 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__6; +x_51 = lean_array_push(x_50, x_49); +lean_ctor_set(x_11, 0, x_51); return x_11; } else { -uint32_t x_51; uint8_t x_52; -x_51 = 10; -lean_inc(x_10); -x_52 = l_String_contains(x_10, x_51); -if (x_52 == 0) +uint32_t x_52; uint8_t x_53; +x_52 = 10; +lean_inc(x_27); +x_53 = l_String_contains(x_27, x_52); +if (x_53 == 0) { -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_53 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__17; -x_54 = lean_string_append(x_53, x_10); -lean_dec(x_10); -x_55 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__18; -x_56 = lean_string_append(x_54, x_55); -x_57 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_57, 0, x_32); -lean_ctor_set(x_57, 1, x_56); -lean_ctor_set(x_57, 2, x_14); -lean_ctor_set(x_57, 3, x_14); -x_58 = l_Lean_Lsp_WorkspaceEdit_ofTextEdit(x_28, x_57); -lean_ctor_set(x_23, 0, x_58); -x_59 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__4; -x_60 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__2; -x_61 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__3; -x_62 = lean_alloc_ctor(0, 10, 0); -lean_ctor_set(x_62, 0, x_14); -lean_ctor_set(x_62, 1, x_14); -lean_ctor_set(x_62, 2, x_59); -lean_ctor_set(x_62, 3, x_60); -lean_ctor_set(x_62, 4, x_14); -lean_ctor_set(x_62, 5, x_61); -lean_ctor_set(x_62, 6, x_14); -lean_ctor_set(x_62, 7, x_23); -lean_ctor_set(x_62, 8, x_14); -lean_ctor_set(x_62, 9, x_14); -x_63 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 1); -lean_closure_set(x_63, 0, x_62); -lean_ctor_set(x_16, 0, x_63); -x_64 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__5; -x_65 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_65, 0, x_64); -lean_ctor_set(x_65, 1, x_16); -x_66 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__6; -x_67 = lean_array_push(x_66, x_65); -lean_ctor_set(x_11, 0, x_67); +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_54 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__17; +x_55 = lean_string_append(x_54, x_27); +lean_dec(x_27); +x_56 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__18; +x_57 = lean_string_append(x_55, x_56); +x_58 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_58, 0, x_33); +lean_ctor_set(x_58, 1, x_57); +lean_ctor_set(x_58, 2, x_14); +lean_ctor_set(x_58, 3, x_14); +x_59 = l_Lean_Lsp_WorkspaceEdit_ofTextEdit(x_29, x_58); +lean_ctor_set(x_23, 0, x_59); +x_60 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__4; +x_61 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__2; +x_62 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__3; +x_63 = lean_alloc_ctor(0, 10, 0); +lean_ctor_set(x_63, 0, x_14); +lean_ctor_set(x_63, 1, x_14); +lean_ctor_set(x_63, 2, x_60); +lean_ctor_set(x_63, 3, x_61); +lean_ctor_set(x_63, 4, x_14); +lean_ctor_set(x_63, 5, x_62); +lean_ctor_set(x_63, 6, x_14); +lean_ctor_set(x_63, 7, x_23); +lean_ctor_set(x_63, 8, x_14); +lean_ctor_set(x_63, 9, x_14); +x_64 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 1); +lean_closure_set(x_64, 0, x_63); +lean_ctor_set(x_16, 0, x_64); +x_65 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__5; +x_66 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_66, 0, x_65); +lean_ctor_set(x_66, 1, x_16); +x_67 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__6; +x_68 = lean_array_push(x_67, x_66); +lean_ctor_set(x_11, 0, x_68); return x_11; } else { -lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; -x_68 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__15; -x_69 = lean_string_append(x_68, x_10); -lean_dec(x_10); -x_70 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__16; -x_71 = lean_string_append(x_69, x_70); -x_72 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_72, 0, x_32); -lean_ctor_set(x_72, 1, x_71); -lean_ctor_set(x_72, 2, x_14); -lean_ctor_set(x_72, 3, x_14); -x_73 = l_Lean_Lsp_WorkspaceEdit_ofTextEdit(x_28, x_72); -lean_ctor_set(x_23, 0, x_73); -x_74 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__4; -x_75 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__2; -x_76 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__3; -x_77 = lean_alloc_ctor(0, 10, 0); -lean_ctor_set(x_77, 0, x_14); -lean_ctor_set(x_77, 1, x_14); -lean_ctor_set(x_77, 2, x_74); -lean_ctor_set(x_77, 3, x_75); -lean_ctor_set(x_77, 4, x_14); -lean_ctor_set(x_77, 5, x_76); -lean_ctor_set(x_77, 6, x_14); -lean_ctor_set(x_77, 7, x_23); -lean_ctor_set(x_77, 8, x_14); -lean_ctor_set(x_77, 9, x_14); -x_78 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 1); -lean_closure_set(x_78, 0, x_77); -lean_ctor_set(x_16, 0, x_78); -x_79 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__5; -x_80 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_80, 0, x_79); -lean_ctor_set(x_80, 1, x_16); -x_81 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__6; -x_82 = lean_array_push(x_81, x_80); -lean_ctor_set(x_11, 0, x_82); +lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; +x_69 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__15; +x_70 = lean_string_append(x_69, x_27); +lean_dec(x_27); +x_71 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__16; +x_72 = lean_string_append(x_70, x_71); +x_73 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_73, 0, x_33); +lean_ctor_set(x_73, 1, x_72); +lean_ctor_set(x_73, 2, x_14); +lean_ctor_set(x_73, 3, x_14); +x_74 = l_Lean_Lsp_WorkspaceEdit_ofTextEdit(x_29, x_73); +lean_ctor_set(x_23, 0, x_74); +x_75 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__4; +x_76 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__2; +x_77 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__3; +x_78 = lean_alloc_ctor(0, 10, 0); +lean_ctor_set(x_78, 0, x_14); +lean_ctor_set(x_78, 1, x_14); +lean_ctor_set(x_78, 2, x_75); +lean_ctor_set(x_78, 3, x_76); +lean_ctor_set(x_78, 4, x_14); +lean_ctor_set(x_78, 5, x_77); +lean_ctor_set(x_78, 6, x_14); +lean_ctor_set(x_78, 7, x_23); +lean_ctor_set(x_78, 8, x_14); +lean_ctor_set(x_78, 9, x_14); +x_79 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 1); +lean_closure_set(x_79, 0, x_78); +lean_ctor_set(x_16, 0, x_79); +x_80 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__5; +x_81 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_81, 0, x_80); +lean_ctor_set(x_81, 1, x_16); +x_82 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__6; +x_83 = lean_array_push(x_82, x_81); +lean_ctor_set(x_11, 0, x_83); return x_11; } } } else { -lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; -lean_dec(x_10); -x_83 = l___private_Lean_Elab_GuardMsgs_0__Lean_Elab_Tactic_GuardMsgs_messageToStringWithoutPos___closed__1; -x_84 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_84, 0, x_32); -lean_ctor_set(x_84, 1, x_83); -lean_ctor_set(x_84, 2, x_14); -lean_ctor_set(x_84, 3, x_14); -x_85 = l_Lean_Lsp_WorkspaceEdit_ofTextEdit(x_28, x_84); -lean_ctor_set(x_23, 0, x_85); -x_86 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__4; -x_87 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__2; -x_88 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__3; -x_89 = lean_alloc_ctor(0, 10, 0); -lean_ctor_set(x_89, 0, x_14); -lean_ctor_set(x_89, 1, x_14); -lean_ctor_set(x_89, 2, x_86); -lean_ctor_set(x_89, 3, x_87); -lean_ctor_set(x_89, 4, x_14); -lean_ctor_set(x_89, 5, x_88); -lean_ctor_set(x_89, 6, x_14); -lean_ctor_set(x_89, 7, x_23); -lean_ctor_set(x_89, 8, x_14); -lean_ctor_set(x_89, 9, x_14); -x_90 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 1); -lean_closure_set(x_90, 0, x_89); -lean_ctor_set(x_16, 0, x_90); -x_91 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__5; -x_92 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_92, 0, x_91); -lean_ctor_set(x_92, 1, x_16); -x_93 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__6; -x_94 = lean_array_push(x_93, x_92); -lean_ctor_set(x_11, 0, x_94); +lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; +lean_dec(x_27); +x_84 = l___private_Lean_Elab_GuardMsgs_0__Lean_Elab_Tactic_GuardMsgs_messageToStringWithoutPos___closed__1; +x_85 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_85, 0, x_33); +lean_ctor_set(x_85, 1, x_84); +lean_ctor_set(x_85, 2, x_14); +lean_ctor_set(x_85, 3, x_14); +x_86 = l_Lean_Lsp_WorkspaceEdit_ofTextEdit(x_29, x_85); +lean_ctor_set(x_23, 0, x_86); +x_87 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__4; +x_88 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__2; +x_89 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__3; +x_90 = lean_alloc_ctor(0, 10, 0); +lean_ctor_set(x_90, 0, x_14); +lean_ctor_set(x_90, 1, x_14); +lean_ctor_set(x_90, 2, x_87); +lean_ctor_set(x_90, 3, x_88); +lean_ctor_set(x_90, 4, x_14); +lean_ctor_set(x_90, 5, x_89); +lean_ctor_set(x_90, 6, x_14); +lean_ctor_set(x_90, 7, x_23); +lean_ctor_set(x_90, 8, x_14); +lean_ctor_set(x_90, 9, x_14); +x_91 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 1); +lean_closure_set(x_91, 0, x_90); +lean_ctor_set(x_16, 0, x_91); +x_92 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__5; +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_92); +lean_ctor_set(x_93, 1, x_16); +x_94 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__6; +x_95 = lean_array_push(x_94, x_93); +lean_ctor_set(x_11, 0, x_95); return x_11; } } else { -lean_object* x_95; uint8_t x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; -x_95 = lean_ctor_get(x_23, 0); -lean_inc(x_95); +lean_object* x_96; lean_object* x_97; uint8_t 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_ctor_get(x_23, 0); +lean_inc(x_96); lean_dec(x_23); -x_96 = l_String_isEmpty(x_10); -x_97 = l_Lean_Server_FileWorker_EditableDocument_versionedIdentifier(x_13); -x_98 = lean_ctor_get(x_13, 0); -lean_inc(x_98); +x_97 = l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace(x_10); +lean_dec(x_10); +x_98 = l_String_isEmpty(x_97); +x_99 = l_Lean_Server_FileWorker_EditableDocument_versionedIdentifier(x_13); +x_100 = lean_ctor_get(x_13, 0); +lean_inc(x_100); lean_dec(x_13); -x_99 = lean_ctor_get(x_98, 2); -lean_inc(x_99); -lean_dec(x_98); -x_100 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_100, 0, x_19); -lean_ctor_set(x_100, 1, x_95); -x_101 = l_Lean_FileMap_utf8RangeToLspRange(x_99, x_100); +x_101 = lean_ctor_get(x_100, 2); +lean_inc(x_101); lean_dec(x_100); -lean_dec(x_99); -if (x_96 == 0) -{ -lean_object* x_102; lean_object* x_103; uint8_t x_104; -x_102 = lean_string_length(x_10); -x_103 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__14; -x_104 = lean_nat_dec_le(x_102, x_103); +x_102 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_102, 0, x_19); +lean_ctor_set(x_102, 1, x_96); +x_103 = l_Lean_FileMap_utf8RangeToLspRange(x_101, x_102); lean_dec(x_102); -if (x_104 == 0) -{ -lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; -x_105 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__15; -x_106 = lean_string_append(x_105, x_10); -lean_dec(x_10); -x_107 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__16; -x_108 = lean_string_append(x_106, x_107); -x_109 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_109, 0, x_101); -lean_ctor_set(x_109, 1, x_108); -lean_ctor_set(x_109, 2, x_14); -lean_ctor_set(x_109, 3, x_14); -x_110 = l_Lean_Lsp_WorkspaceEdit_ofTextEdit(x_97, x_109); -x_111 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_111, 0, x_110); -x_112 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__4; -x_113 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__2; -x_114 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__3; -x_115 = lean_alloc_ctor(0, 10, 0); -lean_ctor_set(x_115, 0, x_14); -lean_ctor_set(x_115, 1, x_14); -lean_ctor_set(x_115, 2, x_112); -lean_ctor_set(x_115, 3, x_113); -lean_ctor_set(x_115, 4, x_14); -lean_ctor_set(x_115, 5, x_114); -lean_ctor_set(x_115, 6, x_14); -lean_ctor_set(x_115, 7, x_111); -lean_ctor_set(x_115, 8, x_14); -lean_ctor_set(x_115, 9, x_14); -x_116 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 1); -lean_closure_set(x_116, 0, x_115); -lean_ctor_set(x_16, 0, x_116); -x_117 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__5; -x_118 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_118, 0, x_117); -lean_ctor_set(x_118, 1, x_16); -x_119 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__6; -x_120 = lean_array_push(x_119, x_118); -lean_ctor_set(x_11, 0, x_120); +lean_dec(x_101); +if (x_98 == 0) +{ +lean_object* x_104; lean_object* x_105; uint8_t x_106; +x_104 = lean_string_length(x_97); +x_105 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__14; +x_106 = lean_nat_dec_le(x_104, x_105); +lean_dec(x_104); +if (x_106 == 0) +{ +lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; +x_107 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__15; +x_108 = lean_string_append(x_107, x_97); +lean_dec(x_97); +x_109 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__16; +x_110 = lean_string_append(x_108, x_109); +x_111 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_111, 0, x_103); +lean_ctor_set(x_111, 1, x_110); +lean_ctor_set(x_111, 2, x_14); +lean_ctor_set(x_111, 3, x_14); +x_112 = l_Lean_Lsp_WorkspaceEdit_ofTextEdit(x_99, x_111); +x_113 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_113, 0, x_112); +x_114 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__4; +x_115 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__2; +x_116 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__3; +x_117 = lean_alloc_ctor(0, 10, 0); +lean_ctor_set(x_117, 0, x_14); +lean_ctor_set(x_117, 1, x_14); +lean_ctor_set(x_117, 2, x_114); +lean_ctor_set(x_117, 3, x_115); +lean_ctor_set(x_117, 4, x_14); +lean_ctor_set(x_117, 5, x_116); +lean_ctor_set(x_117, 6, x_14); +lean_ctor_set(x_117, 7, x_113); +lean_ctor_set(x_117, 8, x_14); +lean_ctor_set(x_117, 9, x_14); +x_118 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 1); +lean_closure_set(x_118, 0, x_117); +lean_ctor_set(x_16, 0, x_118); +x_119 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__5; +x_120 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_120, 0, x_119); +lean_ctor_set(x_120, 1, x_16); +x_121 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__6; +x_122 = lean_array_push(x_121, x_120); +lean_ctor_set(x_11, 0, x_122); return x_11; } else { -uint32_t x_121; uint8_t x_122; -x_121 = 10; -lean_inc(x_10); -x_122 = l_String_contains(x_10, x_121); -if (x_122 == 0) +uint32_t x_123; uint8_t x_124; +x_123 = 10; +lean_inc(x_97); +x_124 = l_String_contains(x_97, x_123); +if (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; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; -x_123 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__17; -x_124 = lean_string_append(x_123, x_10); -lean_dec(x_10); -x_125 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__18; -x_126 = lean_string_append(x_124, x_125); -x_127 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_127, 0, x_101); -lean_ctor_set(x_127, 1, x_126); -lean_ctor_set(x_127, 2, x_14); -lean_ctor_set(x_127, 3, x_14); -x_128 = l_Lean_Lsp_WorkspaceEdit_ofTextEdit(x_97, x_127); -x_129 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_129, 0, x_128); -x_130 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__4; -x_131 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__2; -x_132 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__3; -x_133 = lean_alloc_ctor(0, 10, 0); -lean_ctor_set(x_133, 0, x_14); -lean_ctor_set(x_133, 1, x_14); -lean_ctor_set(x_133, 2, x_130); -lean_ctor_set(x_133, 3, x_131); -lean_ctor_set(x_133, 4, x_14); -lean_ctor_set(x_133, 5, x_132); -lean_ctor_set(x_133, 6, x_14); -lean_ctor_set(x_133, 7, x_129); -lean_ctor_set(x_133, 8, x_14); -lean_ctor_set(x_133, 9, x_14); -x_134 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 1); -lean_closure_set(x_134, 0, x_133); -lean_ctor_set(x_16, 0, x_134); -x_135 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__5; -x_136 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_136, 0, x_135); -lean_ctor_set(x_136, 1, x_16); -x_137 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__6; -x_138 = lean_array_push(x_137, x_136); -lean_ctor_set(x_11, 0, x_138); +lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; +x_125 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__17; +x_126 = lean_string_append(x_125, x_97); +lean_dec(x_97); +x_127 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__18; +x_128 = lean_string_append(x_126, x_127); +x_129 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_129, 0, x_103); +lean_ctor_set(x_129, 1, x_128); +lean_ctor_set(x_129, 2, x_14); +lean_ctor_set(x_129, 3, x_14); +x_130 = l_Lean_Lsp_WorkspaceEdit_ofTextEdit(x_99, x_129); +x_131 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_131, 0, x_130); +x_132 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__4; +x_133 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__2; +x_134 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__3; +x_135 = lean_alloc_ctor(0, 10, 0); +lean_ctor_set(x_135, 0, x_14); +lean_ctor_set(x_135, 1, x_14); +lean_ctor_set(x_135, 2, x_132); +lean_ctor_set(x_135, 3, x_133); +lean_ctor_set(x_135, 4, x_14); +lean_ctor_set(x_135, 5, x_134); +lean_ctor_set(x_135, 6, x_14); +lean_ctor_set(x_135, 7, x_131); +lean_ctor_set(x_135, 8, x_14); +lean_ctor_set(x_135, 9, x_14); +x_136 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 1); +lean_closure_set(x_136, 0, x_135); +lean_ctor_set(x_16, 0, x_136); +x_137 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__5; +x_138 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_138, 0, x_137); +lean_ctor_set(x_138, 1, x_16); +x_139 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__6; +x_140 = lean_array_push(x_139, x_138); +lean_ctor_set(x_11, 0, x_140); return x_11; } else { -lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; -x_139 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__15; -x_140 = lean_string_append(x_139, x_10); -lean_dec(x_10); -x_141 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__16; -x_142 = lean_string_append(x_140, x_141); -x_143 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_143, 0, x_101); -lean_ctor_set(x_143, 1, x_142); -lean_ctor_set(x_143, 2, x_14); -lean_ctor_set(x_143, 3, x_14); -x_144 = l_Lean_Lsp_WorkspaceEdit_ofTextEdit(x_97, x_143); -x_145 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_145, 0, x_144); -x_146 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__4; -x_147 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__2; -x_148 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__3; -x_149 = lean_alloc_ctor(0, 10, 0); -lean_ctor_set(x_149, 0, x_14); -lean_ctor_set(x_149, 1, x_14); -lean_ctor_set(x_149, 2, x_146); -lean_ctor_set(x_149, 3, x_147); -lean_ctor_set(x_149, 4, x_14); -lean_ctor_set(x_149, 5, x_148); -lean_ctor_set(x_149, 6, x_14); -lean_ctor_set(x_149, 7, x_145); -lean_ctor_set(x_149, 8, x_14); -lean_ctor_set(x_149, 9, x_14); -x_150 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 1); -lean_closure_set(x_150, 0, x_149); -lean_ctor_set(x_16, 0, x_150); -x_151 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__5; -x_152 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_152, 0, x_151); -lean_ctor_set(x_152, 1, x_16); -x_153 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__6; -x_154 = lean_array_push(x_153, x_152); -lean_ctor_set(x_11, 0, x_154); +lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; +x_141 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__15; +x_142 = lean_string_append(x_141, x_97); +lean_dec(x_97); +x_143 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__16; +x_144 = lean_string_append(x_142, x_143); +x_145 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_145, 0, x_103); +lean_ctor_set(x_145, 1, x_144); +lean_ctor_set(x_145, 2, x_14); +lean_ctor_set(x_145, 3, x_14); +x_146 = l_Lean_Lsp_WorkspaceEdit_ofTextEdit(x_99, x_145); +x_147 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_147, 0, x_146); +x_148 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__4; +x_149 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__2; +x_150 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__3; +x_151 = lean_alloc_ctor(0, 10, 0); +lean_ctor_set(x_151, 0, x_14); +lean_ctor_set(x_151, 1, x_14); +lean_ctor_set(x_151, 2, x_148); +lean_ctor_set(x_151, 3, x_149); +lean_ctor_set(x_151, 4, x_14); +lean_ctor_set(x_151, 5, x_150); +lean_ctor_set(x_151, 6, x_14); +lean_ctor_set(x_151, 7, x_147); +lean_ctor_set(x_151, 8, x_14); +lean_ctor_set(x_151, 9, x_14); +x_152 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 1); +lean_closure_set(x_152, 0, x_151); +lean_ctor_set(x_16, 0, x_152); +x_153 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__5; +x_154 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_154, 0, x_153); +lean_ctor_set(x_154, 1, x_16); +x_155 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__6; +x_156 = lean_array_push(x_155, x_154); +lean_ctor_set(x_11, 0, x_156); return x_11; } } } else { -lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; -lean_dec(x_10); -x_155 = l___private_Lean_Elab_GuardMsgs_0__Lean_Elab_Tactic_GuardMsgs_messageToStringWithoutPos___closed__1; -x_156 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_156, 0, x_101); -lean_ctor_set(x_156, 1, x_155); -lean_ctor_set(x_156, 2, x_14); -lean_ctor_set(x_156, 3, x_14); -x_157 = l_Lean_Lsp_WorkspaceEdit_ofTextEdit(x_97, x_156); -x_158 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_158, 0, x_157); -x_159 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__4; -x_160 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__2; -x_161 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__3; -x_162 = lean_alloc_ctor(0, 10, 0); -lean_ctor_set(x_162, 0, x_14); -lean_ctor_set(x_162, 1, x_14); -lean_ctor_set(x_162, 2, x_159); -lean_ctor_set(x_162, 3, x_160); -lean_ctor_set(x_162, 4, x_14); -lean_ctor_set(x_162, 5, x_161); -lean_ctor_set(x_162, 6, x_14); -lean_ctor_set(x_162, 7, x_158); -lean_ctor_set(x_162, 8, x_14); -lean_ctor_set(x_162, 9, x_14); -x_163 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 1); -lean_closure_set(x_163, 0, x_162); -lean_ctor_set(x_16, 0, x_163); -x_164 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__5; -x_165 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_165, 0, x_164); -lean_ctor_set(x_165, 1, x_16); -x_166 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__6; -x_167 = lean_array_push(x_166, x_165); -lean_ctor_set(x_11, 0, x_167); +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_dec(x_97); +x_157 = l___private_Lean_Elab_GuardMsgs_0__Lean_Elab_Tactic_GuardMsgs_messageToStringWithoutPos___closed__1; +x_158 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_158, 0, x_103); +lean_ctor_set(x_158, 1, x_157); +lean_ctor_set(x_158, 2, x_14); +lean_ctor_set(x_158, 3, x_14); +x_159 = l_Lean_Lsp_WorkspaceEdit_ofTextEdit(x_99, x_158); +x_160 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_160, 0, x_159); +x_161 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__4; +x_162 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__2; +x_163 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__3; +x_164 = lean_alloc_ctor(0, 10, 0); +lean_ctor_set(x_164, 0, x_14); +lean_ctor_set(x_164, 1, x_14); +lean_ctor_set(x_164, 2, x_161); +lean_ctor_set(x_164, 3, x_162); +lean_ctor_set(x_164, 4, x_14); +lean_ctor_set(x_164, 5, x_163); +lean_ctor_set(x_164, 6, x_14); +lean_ctor_set(x_164, 7, x_160); +lean_ctor_set(x_164, 8, x_14); +lean_ctor_set(x_164, 9, x_14); +x_165 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 1); +lean_closure_set(x_165, 0, x_164); +lean_ctor_set(x_16, 0, x_165); +x_166 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__5; +x_167 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_167, 0, x_166); +lean_ctor_set(x_167, 1, x_16); +x_168 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__6; +x_169 = lean_array_push(x_168, x_167); +lean_ctor_set(x_11, 0, x_169); return x_11; } } @@ -4846,249 +4973,251 @@ return x_11; } else { -lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; -x_168 = lean_ctor_get(x_16, 0); -lean_inc(x_168); +lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; +x_170 = lean_ctor_get(x_16, 0); +lean_inc(x_170); lean_dec(x_16); -x_169 = lean_unsigned_to_nat(0u); -x_170 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__13; -x_171 = l_Lean_Syntax_setArg(x_9, x_169, x_170); -x_172 = l_Lean_Syntax_getPos_x3f(x_171, x_15); -lean_dec(x_171); -if (lean_obj_tag(x_172) == 0) -{ -lean_object* x_173; -lean_dec(x_168); +x_171 = lean_unsigned_to_nat(0u); +x_172 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__13; +x_173 = l_Lean_Syntax_setArg(x_9, x_171, x_172); +x_174 = l_Lean_Syntax_getPos_x3f(x_173, x_15); +lean_dec(x_173); +if (lean_obj_tag(x_174) == 0) +{ +lean_object* x_175; +lean_dec(x_170); lean_dec(x_13); lean_dec(x_10); -x_173 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__10; -lean_ctor_set(x_11, 0, x_173); +x_175 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__10; +lean_ctor_set(x_11, 0, x_175); return x_11; } else { -lean_object* x_174; lean_object* x_175; uint8_t x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; -x_174 = lean_ctor_get(x_172, 0); -lean_inc(x_174); -if (lean_is_exclusive(x_172)) { - lean_ctor_release(x_172, 0); - x_175 = x_172; +lean_object* x_176; lean_object* x_177; lean_object* x_178; uint8_t x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; +x_176 = lean_ctor_get(x_174, 0); +lean_inc(x_176); +if (lean_is_exclusive(x_174)) { + lean_ctor_release(x_174, 0); + x_177 = x_174; } else { - lean_dec_ref(x_172); - x_175 = lean_box(0); + lean_dec_ref(x_174); + x_177 = lean_box(0); } -x_176 = l_String_isEmpty(x_10); -x_177 = l_Lean_Server_FileWorker_EditableDocument_versionedIdentifier(x_13); -x_178 = lean_ctor_get(x_13, 0); -lean_inc(x_178); +x_178 = l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace(x_10); +lean_dec(x_10); +x_179 = l_String_isEmpty(x_178); +x_180 = l_Lean_Server_FileWorker_EditableDocument_versionedIdentifier(x_13); +x_181 = lean_ctor_get(x_13, 0); +lean_inc(x_181); lean_dec(x_13); -x_179 = lean_ctor_get(x_178, 2); -lean_inc(x_179); -lean_dec(x_178); -x_180 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_180, 0, x_168); -lean_ctor_set(x_180, 1, x_174); -x_181 = l_Lean_FileMap_utf8RangeToLspRange(x_179, x_180); -lean_dec(x_180); -lean_dec(x_179); -if (x_176 == 0) -{ -lean_object* x_182; lean_object* x_183; uint8_t x_184; -x_182 = lean_string_length(x_10); -x_183 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__14; -x_184 = lean_nat_dec_le(x_182, x_183); +x_182 = lean_ctor_get(x_181, 2); +lean_inc(x_182); +lean_dec(x_181); +x_183 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_183, 0, x_170); +lean_ctor_set(x_183, 1, x_176); +x_184 = l_Lean_FileMap_utf8RangeToLspRange(x_182, x_183); +lean_dec(x_183); lean_dec(x_182); -if (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; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; -x_185 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__15; -x_186 = lean_string_append(x_185, x_10); -lean_dec(x_10); -x_187 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__16; -x_188 = lean_string_append(x_186, x_187); -x_189 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_189, 0, x_181); -lean_ctor_set(x_189, 1, x_188); -lean_ctor_set(x_189, 2, x_14); -lean_ctor_set(x_189, 3, x_14); -x_190 = l_Lean_Lsp_WorkspaceEdit_ofTextEdit(x_177, x_189); -if (lean_is_scalar(x_175)) { - x_191 = lean_alloc_ctor(1, 1, 0); +if (x_179 == 0) +{ +lean_object* x_185; lean_object* x_186; uint8_t x_187; +x_185 = lean_string_length(x_178); +x_186 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__14; +x_187 = lean_nat_dec_le(x_185, x_186); +lean_dec(x_185); +if (x_187 == 0) +{ +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; +x_188 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__15; +x_189 = lean_string_append(x_188, x_178); +lean_dec(x_178); +x_190 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__16; +x_191 = lean_string_append(x_189, x_190); +x_192 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_192, 0, x_184); +lean_ctor_set(x_192, 1, x_191); +lean_ctor_set(x_192, 2, x_14); +lean_ctor_set(x_192, 3, x_14); +x_193 = l_Lean_Lsp_WorkspaceEdit_ofTextEdit(x_180, x_192); +if (lean_is_scalar(x_177)) { + x_194 = lean_alloc_ctor(1, 1, 0); } else { - x_191 = x_175; -} -lean_ctor_set(x_191, 0, x_190); -x_192 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__4; -x_193 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__2; -x_194 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__3; -x_195 = lean_alloc_ctor(0, 10, 0); -lean_ctor_set(x_195, 0, x_14); -lean_ctor_set(x_195, 1, x_14); -lean_ctor_set(x_195, 2, x_192); -lean_ctor_set(x_195, 3, x_193); -lean_ctor_set(x_195, 4, x_14); -lean_ctor_set(x_195, 5, x_194); -lean_ctor_set(x_195, 6, x_14); -lean_ctor_set(x_195, 7, x_191); -lean_ctor_set(x_195, 8, x_14); -lean_ctor_set(x_195, 9, x_14); -x_196 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 1); -lean_closure_set(x_196, 0, x_195); -x_197 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_197, 0, x_196); -x_198 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__5; -x_199 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_199, 0, x_198); -lean_ctor_set(x_199, 1, x_197); -x_200 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__6; -x_201 = lean_array_push(x_200, x_199); -lean_ctor_set(x_11, 0, x_201); + x_194 = x_177; +} +lean_ctor_set(x_194, 0, x_193); +x_195 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__4; +x_196 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__2; +x_197 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__3; +x_198 = lean_alloc_ctor(0, 10, 0); +lean_ctor_set(x_198, 0, x_14); +lean_ctor_set(x_198, 1, x_14); +lean_ctor_set(x_198, 2, x_195); +lean_ctor_set(x_198, 3, x_196); +lean_ctor_set(x_198, 4, x_14); +lean_ctor_set(x_198, 5, x_197); +lean_ctor_set(x_198, 6, x_14); +lean_ctor_set(x_198, 7, x_194); +lean_ctor_set(x_198, 8, x_14); +lean_ctor_set(x_198, 9, x_14); +x_199 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 1); +lean_closure_set(x_199, 0, x_198); +x_200 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_200, 0, x_199); +x_201 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__5; +x_202 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_202, 0, x_201); +lean_ctor_set(x_202, 1, x_200); +x_203 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__6; +x_204 = lean_array_push(x_203, x_202); +lean_ctor_set(x_11, 0, x_204); return x_11; } else { -uint32_t x_202; uint8_t x_203; -x_202 = 10; -lean_inc(x_10); -x_203 = l_String_contains(x_10, x_202); -if (x_203 == 0) +uint32_t x_205; uint8_t x_206; +x_205 = 10; +lean_inc(x_178); +x_206 = l_String_contains(x_178, x_205); +if (x_206 == 0) { -lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; -x_204 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__17; -x_205 = lean_string_append(x_204, x_10); -lean_dec(x_10); -x_206 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__18; -x_207 = lean_string_append(x_205, x_206); -x_208 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_208, 0, x_181); -lean_ctor_set(x_208, 1, x_207); -lean_ctor_set(x_208, 2, x_14); -lean_ctor_set(x_208, 3, x_14); -x_209 = l_Lean_Lsp_WorkspaceEdit_ofTextEdit(x_177, x_208); -if (lean_is_scalar(x_175)) { - x_210 = lean_alloc_ctor(1, 1, 0); +lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; +x_207 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__17; +x_208 = lean_string_append(x_207, x_178); +lean_dec(x_178); +x_209 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__18; +x_210 = lean_string_append(x_208, x_209); +x_211 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_211, 0, x_184); +lean_ctor_set(x_211, 1, x_210); +lean_ctor_set(x_211, 2, x_14); +lean_ctor_set(x_211, 3, x_14); +x_212 = l_Lean_Lsp_WorkspaceEdit_ofTextEdit(x_180, x_211); +if (lean_is_scalar(x_177)) { + x_213 = lean_alloc_ctor(1, 1, 0); } else { - x_210 = x_175; -} -lean_ctor_set(x_210, 0, x_209); -x_211 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__4; -x_212 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__2; -x_213 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__3; -x_214 = lean_alloc_ctor(0, 10, 0); -lean_ctor_set(x_214, 0, x_14); -lean_ctor_set(x_214, 1, x_14); -lean_ctor_set(x_214, 2, x_211); -lean_ctor_set(x_214, 3, x_212); -lean_ctor_set(x_214, 4, x_14); -lean_ctor_set(x_214, 5, x_213); -lean_ctor_set(x_214, 6, x_14); -lean_ctor_set(x_214, 7, x_210); -lean_ctor_set(x_214, 8, x_14); -lean_ctor_set(x_214, 9, x_14); -x_215 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 1); -lean_closure_set(x_215, 0, x_214); -x_216 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_216, 0, x_215); -x_217 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__5; -x_218 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_218, 0, x_217); -lean_ctor_set(x_218, 1, x_216); -x_219 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__6; -x_220 = lean_array_push(x_219, x_218); -lean_ctor_set(x_11, 0, x_220); + x_213 = x_177; +} +lean_ctor_set(x_213, 0, x_212); +x_214 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__4; +x_215 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__2; +x_216 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__3; +x_217 = lean_alloc_ctor(0, 10, 0); +lean_ctor_set(x_217, 0, x_14); +lean_ctor_set(x_217, 1, x_14); +lean_ctor_set(x_217, 2, x_214); +lean_ctor_set(x_217, 3, x_215); +lean_ctor_set(x_217, 4, x_14); +lean_ctor_set(x_217, 5, x_216); +lean_ctor_set(x_217, 6, x_14); +lean_ctor_set(x_217, 7, x_213); +lean_ctor_set(x_217, 8, x_14); +lean_ctor_set(x_217, 9, x_14); +x_218 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 1); +lean_closure_set(x_218, 0, x_217); +x_219 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_219, 0, x_218); +x_220 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__5; +x_221 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_221, 0, x_220); +lean_ctor_set(x_221, 1, x_219); +x_222 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__6; +x_223 = lean_array_push(x_222, x_221); +lean_ctor_set(x_11, 0, x_223); return x_11; } else { -lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; -x_221 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__15; -x_222 = lean_string_append(x_221, x_10); -lean_dec(x_10); -x_223 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__16; -x_224 = lean_string_append(x_222, x_223); -x_225 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_225, 0, x_181); -lean_ctor_set(x_225, 1, x_224); -lean_ctor_set(x_225, 2, x_14); -lean_ctor_set(x_225, 3, x_14); -x_226 = l_Lean_Lsp_WorkspaceEdit_ofTextEdit(x_177, x_225); -if (lean_is_scalar(x_175)) { - x_227 = lean_alloc_ctor(1, 1, 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; lean_object* x_240; +x_224 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__15; +x_225 = lean_string_append(x_224, x_178); +lean_dec(x_178); +x_226 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__16; +x_227 = lean_string_append(x_225, x_226); +x_228 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_228, 0, x_184); +lean_ctor_set(x_228, 1, x_227); +lean_ctor_set(x_228, 2, x_14); +lean_ctor_set(x_228, 3, x_14); +x_229 = l_Lean_Lsp_WorkspaceEdit_ofTextEdit(x_180, x_228); +if (lean_is_scalar(x_177)) { + x_230 = lean_alloc_ctor(1, 1, 0); } else { - x_227 = x_175; -} -lean_ctor_set(x_227, 0, x_226); -x_228 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__4; -x_229 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__2; -x_230 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__3; -x_231 = lean_alloc_ctor(0, 10, 0); -lean_ctor_set(x_231, 0, x_14); -lean_ctor_set(x_231, 1, x_14); -lean_ctor_set(x_231, 2, x_228); -lean_ctor_set(x_231, 3, x_229); -lean_ctor_set(x_231, 4, x_14); -lean_ctor_set(x_231, 5, x_230); -lean_ctor_set(x_231, 6, x_14); -lean_ctor_set(x_231, 7, x_227); -lean_ctor_set(x_231, 8, x_14); -lean_ctor_set(x_231, 9, x_14); -x_232 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 1); -lean_closure_set(x_232, 0, x_231); -x_233 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_233, 0, x_232); -x_234 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__5; -x_235 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_235, 0, x_234); -lean_ctor_set(x_235, 1, x_233); -x_236 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__6; -x_237 = lean_array_push(x_236, x_235); -lean_ctor_set(x_11, 0, x_237); + x_230 = x_177; +} +lean_ctor_set(x_230, 0, x_229); +x_231 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__4; +x_232 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__2; +x_233 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__3; +x_234 = lean_alloc_ctor(0, 10, 0); +lean_ctor_set(x_234, 0, x_14); +lean_ctor_set(x_234, 1, x_14); +lean_ctor_set(x_234, 2, x_231); +lean_ctor_set(x_234, 3, x_232); +lean_ctor_set(x_234, 4, x_14); +lean_ctor_set(x_234, 5, x_233); +lean_ctor_set(x_234, 6, x_14); +lean_ctor_set(x_234, 7, x_230); +lean_ctor_set(x_234, 8, x_14); +lean_ctor_set(x_234, 9, x_14); +x_235 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 1); +lean_closure_set(x_235, 0, x_234); +x_236 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_236, 0, x_235); +x_237 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__5; +x_238 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_238, 0, x_237); +lean_ctor_set(x_238, 1, x_236); +x_239 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__6; +x_240 = lean_array_push(x_239, x_238); +lean_ctor_set(x_11, 0, x_240); return x_11; } } } else { -lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; -lean_dec(x_10); -x_238 = l___private_Lean_Elab_GuardMsgs_0__Lean_Elab_Tactic_GuardMsgs_messageToStringWithoutPos___closed__1; -x_239 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_239, 0, x_181); -lean_ctor_set(x_239, 1, x_238); -lean_ctor_set(x_239, 2, x_14); -lean_ctor_set(x_239, 3, x_14); -x_240 = l_Lean_Lsp_WorkspaceEdit_ofTextEdit(x_177, x_239); -if (lean_is_scalar(x_175)) { - x_241 = lean_alloc_ctor(1, 1, 0); +lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; +lean_dec(x_178); +x_241 = l___private_Lean_Elab_GuardMsgs_0__Lean_Elab_Tactic_GuardMsgs_messageToStringWithoutPos___closed__1; +x_242 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_242, 0, x_184); +lean_ctor_set(x_242, 1, x_241); +lean_ctor_set(x_242, 2, x_14); +lean_ctor_set(x_242, 3, x_14); +x_243 = l_Lean_Lsp_WorkspaceEdit_ofTextEdit(x_180, x_242); +if (lean_is_scalar(x_177)) { + x_244 = lean_alloc_ctor(1, 1, 0); } else { - x_241 = x_175; -} -lean_ctor_set(x_241, 0, x_240); -x_242 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__4; -x_243 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__2; -x_244 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__3; -x_245 = lean_alloc_ctor(0, 10, 0); -lean_ctor_set(x_245, 0, x_14); -lean_ctor_set(x_245, 1, x_14); -lean_ctor_set(x_245, 2, x_242); -lean_ctor_set(x_245, 3, x_243); -lean_ctor_set(x_245, 4, x_14); -lean_ctor_set(x_245, 5, x_244); -lean_ctor_set(x_245, 6, x_14); -lean_ctor_set(x_245, 7, x_241); -lean_ctor_set(x_245, 8, x_14); -lean_ctor_set(x_245, 9, x_14); -x_246 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 1); -lean_closure_set(x_246, 0, x_245); -x_247 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_247, 0, x_246); -x_248 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__5; -x_249 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_249, 0, x_248); -lean_ctor_set(x_249, 1, x_247); -x_250 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__6; -x_251 = lean_array_push(x_250, x_249); -lean_ctor_set(x_11, 0, x_251); + x_244 = x_177; +} +lean_ctor_set(x_244, 0, x_243); +x_245 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__4; +x_246 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__2; +x_247 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__3; +x_248 = lean_alloc_ctor(0, 10, 0); +lean_ctor_set(x_248, 0, x_14); +lean_ctor_set(x_248, 1, x_14); +lean_ctor_set(x_248, 2, x_245); +lean_ctor_set(x_248, 3, x_246); +lean_ctor_set(x_248, 4, x_14); +lean_ctor_set(x_248, 5, x_247); +lean_ctor_set(x_248, 6, x_14); +lean_ctor_set(x_248, 7, x_244); +lean_ctor_set(x_248, 8, x_14); +lean_ctor_set(x_248, 9, x_14); +x_249 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 1); +lean_closure_set(x_249, 0, x_248); +x_250 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_250, 0, x_249); +x_251 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__5; +x_252 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_252, 0, x_251); +lean_ctor_set(x_252, 1, x_250); +x_253 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__6; +x_254 = lean_array_push(x_253, x_252); +lean_ctor_set(x_11, 0, x_254); return x_11; } } @@ -5097,306 +5226,308 @@ return x_11; } else { -lean_object* x_252; lean_object* x_253; lean_object* x_254; uint8_t x_255; lean_object* x_256; -x_252 = lean_ctor_get(x_11, 0); -x_253 = lean_ctor_get(x_11, 1); -lean_inc(x_253); -lean_inc(x_252); +lean_object* x_255; lean_object* x_256; lean_object* x_257; uint8_t x_258; lean_object* x_259; +x_255 = lean_ctor_get(x_11, 0); +x_256 = lean_ctor_get(x_11, 1); +lean_inc(x_256); +lean_inc(x_255); lean_dec(x_11); -x_254 = lean_box(0); -x_255 = 1; -x_256 = l_Lean_Syntax_getPos_x3f(x_9, x_255); -if (lean_obj_tag(x_256) == 0) +x_257 = lean_box(0); +x_258 = 1; +x_259 = l_Lean_Syntax_getPos_x3f(x_9, x_258); +if (lean_obj_tag(x_259) == 0) { -lean_object* x_257; lean_object* x_258; -lean_dec(x_252); +lean_object* x_260; lean_object* x_261; +lean_dec(x_255); lean_dec(x_10); lean_dec(x_9); -x_257 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__10; -x_258 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_258, 0, x_257); -lean_ctor_set(x_258, 1, x_253); -return x_258; +x_260 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__10; +x_261 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_261, 0, x_260); +lean_ctor_set(x_261, 1, x_256); +return x_261; } else { -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_259 = lean_ctor_get(x_256, 0); -lean_inc(x_259); -if (lean_is_exclusive(x_256)) { - lean_ctor_release(x_256, 0); - x_260 = x_256; +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_262 = lean_ctor_get(x_259, 0); +lean_inc(x_262); +if (lean_is_exclusive(x_259)) { + lean_ctor_release(x_259, 0); + x_263 = x_259; } else { - lean_dec_ref(x_256); - x_260 = lean_box(0); + lean_dec_ref(x_259); + x_263 = lean_box(0); } -x_261 = lean_unsigned_to_nat(0u); -x_262 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__13; -x_263 = l_Lean_Syntax_setArg(x_9, x_261, x_262); -x_264 = l_Lean_Syntax_getPos_x3f(x_263, x_255); -lean_dec(x_263); -if (lean_obj_tag(x_264) == 0) +x_264 = lean_unsigned_to_nat(0u); +x_265 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__13; +x_266 = l_Lean_Syntax_setArg(x_9, x_264, x_265); +x_267 = l_Lean_Syntax_getPos_x3f(x_266, x_258); +lean_dec(x_266); +if (lean_obj_tag(x_267) == 0) { -lean_object* x_265; lean_object* x_266; -lean_dec(x_260); -lean_dec(x_259); -lean_dec(x_252); +lean_object* x_268; lean_object* x_269; +lean_dec(x_263); +lean_dec(x_262); +lean_dec(x_255); lean_dec(x_10); -x_265 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__10; -x_266 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_266, 0, x_265); -lean_ctor_set(x_266, 1, x_253); -return x_266; +x_268 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__10; +x_269 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_269, 0, x_268); +lean_ctor_set(x_269, 1, x_256); +return x_269; } else { -lean_object* x_267; lean_object* x_268; uint8_t x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; -x_267 = lean_ctor_get(x_264, 0); -lean_inc(x_267); -if (lean_is_exclusive(x_264)) { - lean_ctor_release(x_264, 0); - x_268 = x_264; +lean_object* x_270; lean_object* x_271; lean_object* x_272; uint8_t x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; +x_270 = lean_ctor_get(x_267, 0); +lean_inc(x_270); +if (lean_is_exclusive(x_267)) { + lean_ctor_release(x_267, 0); + x_271 = x_267; } else { - lean_dec_ref(x_264); - x_268 = lean_box(0); -} -x_269 = l_String_isEmpty(x_10); -x_270 = l_Lean_Server_FileWorker_EditableDocument_versionedIdentifier(x_252); -x_271 = lean_ctor_get(x_252, 0); -lean_inc(x_271); -lean_dec(x_252); -x_272 = lean_ctor_get(x_271, 2); -lean_inc(x_272); -lean_dec(x_271); -x_273 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_273, 0, x_259); -lean_ctor_set(x_273, 1, x_267); -x_274 = l_Lean_FileMap_utf8RangeToLspRange(x_272, x_273); -lean_dec(x_273); -lean_dec(x_272); -if (x_269 == 0) -{ -lean_object* x_275; lean_object* x_276; uint8_t x_277; -x_275 = lean_string_length(x_10); -x_276 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__14; -x_277 = lean_nat_dec_le(x_275, x_276); -lean_dec(x_275); -if (x_277 == 0) -{ -lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_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; -x_278 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__15; -x_279 = lean_string_append(x_278, x_10); + lean_dec_ref(x_267); + x_271 = lean_box(0); +} +x_272 = l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace(x_10); lean_dec(x_10); -x_280 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__16; -x_281 = lean_string_append(x_279, x_280); -x_282 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_282, 0, x_274); -lean_ctor_set(x_282, 1, x_281); -lean_ctor_set(x_282, 2, x_254); -lean_ctor_set(x_282, 3, x_254); -x_283 = l_Lean_Lsp_WorkspaceEdit_ofTextEdit(x_270, x_282); -if (lean_is_scalar(x_268)) { - x_284 = lean_alloc_ctor(1, 1, 0); +x_273 = l_String_isEmpty(x_272); +x_274 = l_Lean_Server_FileWorker_EditableDocument_versionedIdentifier(x_255); +x_275 = lean_ctor_get(x_255, 0); +lean_inc(x_275); +lean_dec(x_255); +x_276 = lean_ctor_get(x_275, 2); +lean_inc(x_276); +lean_dec(x_275); +x_277 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_277, 0, x_262); +lean_ctor_set(x_277, 1, x_270); +x_278 = l_Lean_FileMap_utf8RangeToLspRange(x_276, x_277); +lean_dec(x_277); +lean_dec(x_276); +if (x_273 == 0) +{ +lean_object* x_279; lean_object* x_280; uint8_t x_281; +x_279 = lean_string_length(x_272); +x_280 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__14; +x_281 = lean_nat_dec_le(x_279, x_280); +lean_dec(x_279); +if (x_281 == 0) +{ +lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; 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_282 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__15; +x_283 = lean_string_append(x_282, x_272); +lean_dec(x_272); +x_284 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__16; +x_285 = lean_string_append(x_283, x_284); +x_286 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_286, 0, x_278); +lean_ctor_set(x_286, 1, x_285); +lean_ctor_set(x_286, 2, x_257); +lean_ctor_set(x_286, 3, x_257); +x_287 = l_Lean_Lsp_WorkspaceEdit_ofTextEdit(x_274, x_286); +if (lean_is_scalar(x_271)) { + x_288 = lean_alloc_ctor(1, 1, 0); } else { - x_284 = x_268; -} -lean_ctor_set(x_284, 0, x_283); -x_285 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__4; -x_286 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__2; -x_287 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__3; -x_288 = lean_alloc_ctor(0, 10, 0); -lean_ctor_set(x_288, 0, x_254); -lean_ctor_set(x_288, 1, x_254); -lean_ctor_set(x_288, 2, x_285); -lean_ctor_set(x_288, 3, x_286); -lean_ctor_set(x_288, 4, x_254); -lean_ctor_set(x_288, 5, x_287); -lean_ctor_set(x_288, 6, x_254); -lean_ctor_set(x_288, 7, x_284); -lean_ctor_set(x_288, 8, x_254); -lean_ctor_set(x_288, 9, x_254); -x_289 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 1); -lean_closure_set(x_289, 0, x_288); -if (lean_is_scalar(x_260)) { - x_290 = lean_alloc_ctor(1, 1, 0); + x_288 = x_271; +} +lean_ctor_set(x_288, 0, x_287); +x_289 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__4; +x_290 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__2; +x_291 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__3; +x_292 = lean_alloc_ctor(0, 10, 0); +lean_ctor_set(x_292, 0, x_257); +lean_ctor_set(x_292, 1, x_257); +lean_ctor_set(x_292, 2, x_289); +lean_ctor_set(x_292, 3, x_290); +lean_ctor_set(x_292, 4, x_257); +lean_ctor_set(x_292, 5, x_291); +lean_ctor_set(x_292, 6, x_257); +lean_ctor_set(x_292, 7, x_288); +lean_ctor_set(x_292, 8, x_257); +lean_ctor_set(x_292, 9, x_257); +x_293 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 1); +lean_closure_set(x_293, 0, x_292); +if (lean_is_scalar(x_263)) { + x_294 = lean_alloc_ctor(1, 1, 0); } else { - x_290 = x_260; -} -lean_ctor_set(x_290, 0, x_289); -x_291 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__5; -x_292 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_292, 0, x_291); -lean_ctor_set(x_292, 1, x_290); -x_293 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__6; -x_294 = lean_array_push(x_293, x_292); -x_295 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_295, 0, x_294); -lean_ctor_set(x_295, 1, x_253); -return x_295; + x_294 = x_263; +} +lean_ctor_set(x_294, 0, x_293); +x_295 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__5; +x_296 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_296, 0, x_295); +lean_ctor_set(x_296, 1, x_294); +x_297 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__6; +x_298 = lean_array_push(x_297, x_296); +x_299 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_299, 0, x_298); +lean_ctor_set(x_299, 1, x_256); +return x_299; } else { -uint32_t x_296; uint8_t x_297; -x_296 = 10; -lean_inc(x_10); -x_297 = l_String_contains(x_10, x_296); -if (x_297 == 0) +uint32_t x_300; uint8_t x_301; +x_300 = 10; +lean_inc(x_272); +x_301 = l_String_contains(x_272, x_300); +if (x_301 == 0) { -lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; -x_298 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__17; -x_299 = lean_string_append(x_298, x_10); -lean_dec(x_10); -x_300 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__18; -x_301 = lean_string_append(x_299, x_300); -x_302 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_302, 0, x_274); -lean_ctor_set(x_302, 1, x_301); -lean_ctor_set(x_302, 2, x_254); -lean_ctor_set(x_302, 3, x_254); -x_303 = l_Lean_Lsp_WorkspaceEdit_ofTextEdit(x_270, x_302); -if (lean_is_scalar(x_268)) { - x_304 = lean_alloc_ctor(1, 1, 0); +lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; +x_302 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__17; +x_303 = lean_string_append(x_302, x_272); +lean_dec(x_272); +x_304 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__18; +x_305 = lean_string_append(x_303, x_304); +x_306 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_306, 0, x_278); +lean_ctor_set(x_306, 1, x_305); +lean_ctor_set(x_306, 2, x_257); +lean_ctor_set(x_306, 3, x_257); +x_307 = l_Lean_Lsp_WorkspaceEdit_ofTextEdit(x_274, x_306); +if (lean_is_scalar(x_271)) { + x_308 = lean_alloc_ctor(1, 1, 0); } else { - x_304 = x_268; -} -lean_ctor_set(x_304, 0, x_303); -x_305 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__4; -x_306 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__2; -x_307 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__3; -x_308 = lean_alloc_ctor(0, 10, 0); -lean_ctor_set(x_308, 0, x_254); -lean_ctor_set(x_308, 1, x_254); -lean_ctor_set(x_308, 2, x_305); -lean_ctor_set(x_308, 3, x_306); -lean_ctor_set(x_308, 4, x_254); -lean_ctor_set(x_308, 5, x_307); -lean_ctor_set(x_308, 6, x_254); -lean_ctor_set(x_308, 7, x_304); -lean_ctor_set(x_308, 8, x_254); -lean_ctor_set(x_308, 9, x_254); -x_309 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 1); -lean_closure_set(x_309, 0, x_308); -if (lean_is_scalar(x_260)) { - x_310 = lean_alloc_ctor(1, 1, 0); + x_308 = x_271; +} +lean_ctor_set(x_308, 0, x_307); +x_309 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__4; +x_310 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__2; +x_311 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__3; +x_312 = lean_alloc_ctor(0, 10, 0); +lean_ctor_set(x_312, 0, x_257); +lean_ctor_set(x_312, 1, x_257); +lean_ctor_set(x_312, 2, x_309); +lean_ctor_set(x_312, 3, x_310); +lean_ctor_set(x_312, 4, x_257); +lean_ctor_set(x_312, 5, x_311); +lean_ctor_set(x_312, 6, x_257); +lean_ctor_set(x_312, 7, x_308); +lean_ctor_set(x_312, 8, x_257); +lean_ctor_set(x_312, 9, x_257); +x_313 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 1); +lean_closure_set(x_313, 0, x_312); +if (lean_is_scalar(x_263)) { + x_314 = lean_alloc_ctor(1, 1, 0); } else { - x_310 = x_260; -} -lean_ctor_set(x_310, 0, x_309); -x_311 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__5; -x_312 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_312, 0, x_311); -lean_ctor_set(x_312, 1, x_310); -x_313 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__6; -x_314 = lean_array_push(x_313, x_312); -x_315 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_315, 0, x_314); -lean_ctor_set(x_315, 1, x_253); -return x_315; + x_314 = x_263; +} +lean_ctor_set(x_314, 0, x_313); +x_315 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__5; +x_316 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_316, 0, x_315); +lean_ctor_set(x_316, 1, x_314); +x_317 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__6; +x_318 = lean_array_push(x_317, x_316); +x_319 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_319, 0, x_318); +lean_ctor_set(x_319, 1, x_256); +return x_319; } else { -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; -x_316 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__15; -x_317 = lean_string_append(x_316, x_10); -lean_dec(x_10); -x_318 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__16; -x_319 = lean_string_append(x_317, x_318); -x_320 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_320, 0, x_274); -lean_ctor_set(x_320, 1, x_319); -lean_ctor_set(x_320, 2, x_254); -lean_ctor_set(x_320, 3, x_254); -x_321 = l_Lean_Lsp_WorkspaceEdit_ofTextEdit(x_270, x_320); -if (lean_is_scalar(x_268)) { - x_322 = lean_alloc_ctor(1, 1, 0); +lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_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; +x_320 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__15; +x_321 = lean_string_append(x_320, x_272); +lean_dec(x_272); +x_322 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__16; +x_323 = lean_string_append(x_321, x_322); +x_324 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_324, 0, x_278); +lean_ctor_set(x_324, 1, x_323); +lean_ctor_set(x_324, 2, x_257); +lean_ctor_set(x_324, 3, x_257); +x_325 = l_Lean_Lsp_WorkspaceEdit_ofTextEdit(x_274, x_324); +if (lean_is_scalar(x_271)) { + x_326 = lean_alloc_ctor(1, 1, 0); } else { - x_322 = x_268; -} -lean_ctor_set(x_322, 0, x_321); -x_323 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__4; -x_324 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__2; -x_325 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__3; -x_326 = lean_alloc_ctor(0, 10, 0); -lean_ctor_set(x_326, 0, x_254); -lean_ctor_set(x_326, 1, x_254); -lean_ctor_set(x_326, 2, x_323); -lean_ctor_set(x_326, 3, x_324); -lean_ctor_set(x_326, 4, x_254); -lean_ctor_set(x_326, 5, x_325); -lean_ctor_set(x_326, 6, x_254); -lean_ctor_set(x_326, 7, x_322); -lean_ctor_set(x_326, 8, x_254); -lean_ctor_set(x_326, 9, x_254); -x_327 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 1); -lean_closure_set(x_327, 0, x_326); -if (lean_is_scalar(x_260)) { - x_328 = lean_alloc_ctor(1, 1, 0); + x_326 = x_271; +} +lean_ctor_set(x_326, 0, x_325); +x_327 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__4; +x_328 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__2; +x_329 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__3; +x_330 = lean_alloc_ctor(0, 10, 0); +lean_ctor_set(x_330, 0, x_257); +lean_ctor_set(x_330, 1, x_257); +lean_ctor_set(x_330, 2, x_327); +lean_ctor_set(x_330, 3, x_328); +lean_ctor_set(x_330, 4, x_257); +lean_ctor_set(x_330, 5, x_329); +lean_ctor_set(x_330, 6, x_257); +lean_ctor_set(x_330, 7, x_326); +lean_ctor_set(x_330, 8, x_257); +lean_ctor_set(x_330, 9, x_257); +x_331 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 1); +lean_closure_set(x_331, 0, x_330); +if (lean_is_scalar(x_263)) { + x_332 = lean_alloc_ctor(1, 1, 0); } else { - x_328 = x_260; -} -lean_ctor_set(x_328, 0, x_327); -x_329 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__5; -x_330 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_330, 0, x_329); -lean_ctor_set(x_330, 1, x_328); -x_331 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__6; -x_332 = lean_array_push(x_331, x_330); -x_333 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_333, 0, x_332); -lean_ctor_set(x_333, 1, x_253); -return x_333; + x_332 = x_263; +} +lean_ctor_set(x_332, 0, x_331); +x_333 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__5; +x_334 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_334, 0, x_333); +lean_ctor_set(x_334, 1, x_332); +x_335 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__6; +x_336 = lean_array_push(x_335, x_334); +x_337 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_337, 0, x_336); +lean_ctor_set(x_337, 1, x_256); +return x_337; } } } else { -lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; -lean_dec(x_10); -x_334 = l___private_Lean_Elab_GuardMsgs_0__Lean_Elab_Tactic_GuardMsgs_messageToStringWithoutPos___closed__1; -x_335 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_335, 0, x_274); -lean_ctor_set(x_335, 1, x_334); -lean_ctor_set(x_335, 2, x_254); -lean_ctor_set(x_335, 3, x_254); -x_336 = l_Lean_Lsp_WorkspaceEdit_ofTextEdit(x_270, x_335); -if (lean_is_scalar(x_268)) { - x_337 = lean_alloc_ctor(1, 1, 0); +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_dec(x_272); +x_338 = l___private_Lean_Elab_GuardMsgs_0__Lean_Elab_Tactic_GuardMsgs_messageToStringWithoutPos___closed__1; +x_339 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_339, 0, x_278); +lean_ctor_set(x_339, 1, x_338); +lean_ctor_set(x_339, 2, x_257); +lean_ctor_set(x_339, 3, x_257); +x_340 = l_Lean_Lsp_WorkspaceEdit_ofTextEdit(x_274, x_339); +if (lean_is_scalar(x_271)) { + x_341 = lean_alloc_ctor(1, 1, 0); } else { - x_337 = x_268; -} -lean_ctor_set(x_337, 0, x_336); -x_338 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__4; -x_339 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__2; -x_340 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__3; -x_341 = lean_alloc_ctor(0, 10, 0); -lean_ctor_set(x_341, 0, x_254); -lean_ctor_set(x_341, 1, x_254); -lean_ctor_set(x_341, 2, x_338); -lean_ctor_set(x_341, 3, x_339); -lean_ctor_set(x_341, 4, x_254); -lean_ctor_set(x_341, 5, x_340); -lean_ctor_set(x_341, 6, x_254); -lean_ctor_set(x_341, 7, x_337); -lean_ctor_set(x_341, 8, x_254); -lean_ctor_set(x_341, 9, x_254); -x_342 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 1); -lean_closure_set(x_342, 0, x_341); -if (lean_is_scalar(x_260)) { - x_343 = lean_alloc_ctor(1, 1, 0); + x_341 = x_271; +} +lean_ctor_set(x_341, 0, x_340); +x_342 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__4; +x_343 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__2; +x_344 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__3; +x_345 = lean_alloc_ctor(0, 10, 0); +lean_ctor_set(x_345, 0, x_257); +lean_ctor_set(x_345, 1, x_257); +lean_ctor_set(x_345, 2, x_342); +lean_ctor_set(x_345, 3, x_343); +lean_ctor_set(x_345, 4, x_257); +lean_ctor_set(x_345, 5, x_344); +lean_ctor_set(x_345, 6, x_257); +lean_ctor_set(x_345, 7, x_341); +lean_ctor_set(x_345, 8, x_257); +lean_ctor_set(x_345, 9, x_257); +x_346 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 1); +lean_closure_set(x_346, 0, x_345); +if (lean_is_scalar(x_263)) { + x_347 = lean_alloc_ctor(1, 1, 0); } else { - x_343 = x_260; + x_347 = x_263; } -lean_ctor_set(x_343, 0, x_342); -x_344 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__5; -x_345 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_345, 0, x_344); -lean_ctor_set(x_345, 1, x_343); -x_346 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__6; -x_347 = lean_array_push(x_346, x_345); -x_348 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_348, 0, x_347); -lean_ctor_set(x_348, 1, x_253); -return x_348; +lean_ctor_set(x_347, 0, x_346); +x_348 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__5; +x_349 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_349, 0, x_348); +lean_ctor_set(x_349, 1, x_347); +x_350 = l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__6; +x_351 = lean_array_push(x_350, x_349); +x_352 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_352, 0, x_351); +lean_ctor_set(x_352, 1, x_256); +return x_352; } } } @@ -5405,12 +5536,12 @@ return x_348; } else { -lean_object* x_349; lean_object* x_350; -x_349 = l_Lean_Elab_Tactic_GuardMsgs_parseGuardMsgsSpec___closed__6; -x_350 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_350, 0, x_349); -lean_ctor_set(x_350, 1, x_3); -return x_350; +lean_object* x_353; lean_object* x_354; +x_353 = l_Lean_Elab_Tactic_GuardMsgs_parseGuardMsgsSpec___closed__6; +x_354 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_354, 0, x_353); +lean_ctor_set(x_354, 1, x_3); +return x_354; } } } @@ -5515,7 +5646,7 @@ lean_dec(x_1); return x_4; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare____x40_Lean_Elab_GuardMsgs___hyg_2251____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare____x40_Lean_Elab_GuardMsgs___hyg_2285____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -5525,7 +5656,7 @@ x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare____x40_Lean_Elab_GuardMsgs___hyg_2251____closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare____x40_Lean_Elab_GuardMsgs___hyg_2285____closed__2() { _start: { lean_object* x_1; @@ -5533,12 +5664,12 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeActio return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare____x40_Lean_Elab_GuardMsgs___hyg_2251_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare____x40_Lean_Elab_GuardMsgs___hyg_2285_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare____x40_Lean_Elab_GuardMsgs___hyg_2251____closed__1; -x_3 = l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare____x40_Lean_Elab_GuardMsgs___hyg_2251____closed__2; +x_2 = l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare____x40_Lean_Elab_GuardMsgs___hyg_2285____closed__1; +x_3 = l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare____x40_Lean_Elab_GuardMsgs___hyg_2285____closed__2; x_4 = l_Lean_CodeAction_insertBuiltin(x_2, x_3, x_1); return x_4; } @@ -5632,6 +5763,18 @@ l_Lean_Elab_Tactic_GuardMsgs_instImpl____x40_Lean_Elab_GuardMsgs___hyg_1100_ = _ lean_mark_persistent(l_Lean_Elab_Tactic_GuardMsgs_instImpl____x40_Lean_Elab_GuardMsgs___hyg_1100_); l_Lean_Elab_Tactic_GuardMsgs_instTypeNameGuardMsgFailure = _init_l_Lean_Elab_Tactic_GuardMsgs_instTypeNameGuardMsgFailure(); lean_mark_persistent(l_Lean_Elab_Tactic_GuardMsgs_instTypeNameGuardMsgFailure); +l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace___closed__1 = _init_l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace___closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace___closed__1); +l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace___closed__2 = _init_l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace___closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace___closed__2); +l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace___closed__3 = _init_l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace___closed__3(); +lean_mark_persistent(l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace___closed__3); +l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace___closed__4 = _init_l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace___closed__4(); +lean_mark_persistent(l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace___closed__4); +l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace___closed__5 = _init_l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace___closed__5(); +lean_mark_persistent(l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace___closed__5); +l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace___closed__6 = _init_l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace___closed__6(); +lean_mark_persistent(l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace___closed__6); l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___spec__1___rarg___closed__1 = _init_l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___spec__1___rarg___closed__1(); lean_mark_persistent(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___spec__1___rarg___closed__1); l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___spec__1___rarg___closed__2 = _init_l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___spec__1___rarg___closed__2(); @@ -5738,11 +5881,11 @@ l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__17 = _init_l_L lean_mark_persistent(l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__17); l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__18 = _init_l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__18(); lean_mark_persistent(l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__18); -l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare____x40_Lean_Elab_GuardMsgs___hyg_2251____closed__1 = _init_l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare____x40_Lean_Elab_GuardMsgs___hyg_2251____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare____x40_Lean_Elab_GuardMsgs___hyg_2251____closed__1); -l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare____x40_Lean_Elab_GuardMsgs___hyg_2251____closed__2 = _init_l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare____x40_Lean_Elab_GuardMsgs___hyg_2251____closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare____x40_Lean_Elab_GuardMsgs___hyg_2251____closed__2); -if (builtin) {res = l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare____x40_Lean_Elab_GuardMsgs___hyg_2251_(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare____x40_Lean_Elab_GuardMsgs___hyg_2285____closed__1 = _init_l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare____x40_Lean_Elab_GuardMsgs___hyg_2285____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare____x40_Lean_Elab_GuardMsgs___hyg_2285____closed__1); +l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare____x40_Lean_Elab_GuardMsgs___hyg_2285____closed__2 = _init_l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare____x40_Lean_Elab_GuardMsgs___hyg_2285____closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare____x40_Lean_Elab_GuardMsgs___hyg_2285____closed__2); +if (builtin) {res = l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare____x40_Lean_Elab_GuardMsgs___hyg_2285_(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/MutualDef.c b/stage0/stdlib/Lean/Elab/MutualDef.c index c602d3693d9b..c5d22ff09d90 100644 --- a/stage0/stdlib/Lean/Elab/MutualDef.c +++ b/stage0/stdlib/Lean/Elab/MutualDef.c @@ -58,7 +58,6 @@ lean_object* l_Lean_Elab_Term_processDefDeriving(lean_object*, lean_object*, lea static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__2___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_instFVarIdSetInhabited; -LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___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_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__4___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabMutualDef___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withFunLocalDecls___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -95,6 +94,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualD 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; +static lean_object* l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___spec__1___closed__1; lean_object* l_Lean_throwError___at_Lean_Elab_Term_mkAuxName___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__4___rarg(lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, size_t, size_t, lean_object*); @@ -103,6 +103,7 @@ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_ 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* l_Lean_Meta_isProp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_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*); @@ -142,6 +143,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDe LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getId(lean_object*); lean_object* l_Lean_Elab_WF_TerminationHints_rememberExtraParams(lean_object*, lean_object*, lean_object*); +static lean_object* l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___spec__1___closed__4; lean_object* l_Lean_Elab_Term_elabTermEnsuringType(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_sort___override(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabModifiers___at_Lean_Elab_Command_elabMutualDef___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -305,7 +307,6 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Mutua LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_updateUsedVarsOf___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_MutualClosure_FixPoint_run___boxed(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_processDeriving___spec__1___closed__1; -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; 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*); @@ -324,7 +325,6 @@ lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Elab_Term_elabMutualDef_go___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___lambda__3___closed__3; LEAN_EXPORT lean_object* l_Lean_RBNode_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_pushNewVars___spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__7___closed__2; uint8_t l_List_beq___at___private_Lean_Declaration_0__Lean_beqConstantVal____x40_Lean_Declaration___hyg_236____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__9; LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_collectUsed___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -567,6 +567,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check static lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Term_MutualClosure_Replacement_apply___boxed(lean_object*, lean_object*); lean_object* l_Lean_Syntax_node2(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Elab_Term_MutualClosure_pushLetRecs___spec__1___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getFunName(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__3; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_MutualClosure_main___spec__2(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -682,6 +683,7 @@ 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; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_checkForHiddenUnivLevels_visit___spec__4___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_EXPORT lean_object* l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___spec__1___lambda__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Parser_Term_type; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3___closed__2; static lean_object* l_Lean_Elab_elabModifiers___at_Lean_Elab_Command_elabMutualDef___spec__1___closed__2; @@ -722,7 +724,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_WF_elabTerminationHints___at___private_Lean LEAN_EXPORT lean_object* l_Lean_RBNode_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); 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*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getFunName___lambda__1___boxed(lean_object*, lean_object*); -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___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_typeHasRecFun___lambda__2___boxed(lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); lean_object* l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_withDeclName___spec__3___rarg(lean_object*, lean_object*); @@ -766,6 +767,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels_visit___lambd static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__5___closed__3; LEAN_EXPORT lean_object* l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___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* l_Lean_Option_set___at_Lean_Elab_Eqns_tryURefl___spec__1(lean_object*, lean_object*, uint8_t); +LEAN_EXPORT lean_object* l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_node1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Option_setIfNotSet___at_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___spec__1___boxed(lean_object*, lean_object*, lean_object*); @@ -790,6 +792,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualD LEAN_EXPORT lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_merge___closed__1; +static lean_object* l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___spec__1___closed__2; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosureFor___lambda__2___closed__1; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__4; lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -887,7 +890,7 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Mutua static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_cleanupOfNat___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap___spec__5___rarg___lambda__1(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Elab_Term_MutualClosure_pushLetRecs___spec__1___lambda__2(lean_object*, uint8_t, 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__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_withLevelNames___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_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___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_Lean_Elab_Term_MutualClosure_FixPoint_State_usedFVarsMap___default; @@ -931,7 +934,6 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Mutua lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__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_getAttributeImpl(lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__7___closed__1; 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*); extern lean_object* l_Lean_pp_letVarTypes; @@ -971,6 +973,7 @@ static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0_ lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_processDeriving___spec__1___closed__4; +static lean_object* l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___spec__1___closed__3; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMutualDef___spec__12___closed__2; uint8_t l_Lean_Exception_isRuntime(lean_object*); lean_object* l_Lean_Expr_lam___override(lean_object*, lean_object*, lean_object*, uint8_t); @@ -1022,6 +1025,7 @@ lean_object* l_Lean_Expr_collectFVars(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*); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Elab_Term_MutualClosure_Replacement_apply___spec__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Elab_Term_MutualClosure_pushLetRecs___spec__1___lambda__3___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_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*); lean_object* l_Lean_Elab_Term_elabType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -2097,7 +2101,7 @@ static lean_object* _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_chec _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("'noncomputable unsafe' is not allowed", 37); +x_1 = lean_mk_string_from_bytes("'theorem' subsumes 'noncomputable', code is not generated for theorems", 70); return x_1; } } @@ -2113,80 +2117,6 @@ return x_2; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_11; uint8_t x_12; -lean_dec(x_3); -x_11 = lean_ctor_get(x_2, 1); -x_12 = lean_ctor_get_uint8(x_11, sizeof(void*)*2 + 1); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -x_13 = lean_box(0); -x_14 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__4(x_1, x_2, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_14; -} -else -{ -uint8_t x_15; -x_15 = lean_ctor_get_uint8(x_11, sizeof(void*)*2 + 3); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; -x_16 = lean_box(0); -x_17 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__4(x_1, x_2, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_17; -} -else -{ -lean_object* x_18; lean_object* x_19; uint8_t x_20; -x_18 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__5___closed__2; -x_19 = l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3(x_18, 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); -x_20 = !lean_is_exclusive(x_19); -if (x_20 == 0) -{ -return x_19; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_19, 0); -x_22 = lean_ctor_get(x_19, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_19); -x_23 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_23, 0, x_21); -lean_ctor_set(x_23, 1, x_22); -return x_23; -} -} -} -} -} -static lean_object* _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__6___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("'theorem' subsumes 'noncomputable', code is not generated for theorems", 70); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__6___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__6___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___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) { -_start: -{ uint8_t x_11; uint8_t x_12; lean_dec(x_3); x_11 = lean_ctor_get_uint8(x_2, sizeof(void*)*9); @@ -2195,7 +2125,7 @@ if (x_12 == 0) { lean_object* x_13; lean_object* x_14; x_13 = lean_box(0); -x_14 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__5(x_1, x_2, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_14 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__4(x_1, x_2, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10); return x_14; } else @@ -2207,13 +2137,13 @@ if (x_16 == 0) { lean_object* x_17; lean_object* x_18; x_17 = lean_box(0); -x_18 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__5(x_1, x_2, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_18 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__4(x_1, x_2, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_10); return x_18; } else { lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_19 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__6___closed__2; +x_19 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__5___closed__2; x_20 = l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3(x_19, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_9); lean_dec(x_8); @@ -2242,7 +2172,7 @@ return x_24; } } } -static lean_object* _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__7___closed__1() { +static lean_object* _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__6___closed__1() { _start: { lean_object* x_1; @@ -2250,16 +2180,16 @@ x_1 = lean_mk_string_from_bytes("'partial' theorems are not allowed, 'partial' i return x_1; } } -static lean_object* _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__7___closed__2() { +static lean_object* _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__6___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__7___closed__1; +x_1 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__6___closed__1; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___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___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___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) { _start: { uint8_t x_11; uint8_t x_12; @@ -2270,7 +2200,7 @@ if (x_12 == 0) { lean_object* x_13; lean_object* x_14; x_13 = lean_box(0); -x_14 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__6(x_1, x_2, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_14 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__5(x_1, x_2, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10); return x_14; } else @@ -2282,13 +2212,13 @@ if (x_16 == 0) { lean_object* x_17; lean_object* x_18; x_17 = lean_box(0); -x_18 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__6(x_1, x_2, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_18 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__5(x_1, x_2, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_10); return x_18; } else { lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_19 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__7___closed__2; +x_19 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__6___closed__2; x_20 = l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3(x_19, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_9); lean_dec(x_8); @@ -2344,7 +2274,7 @@ if (x_11 == 0) { lean_object* x_12; lean_object* x_13; x_12 = lean_box(0); -x_13 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__7(x_1, x_2, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_13 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__6(x_1, x_2, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_13; } else @@ -2356,7 +2286,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; x_16 = lean_box(0); -x_17 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__7(x_1, x_2, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_17 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__6(x_1, x_2, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_17; } else @@ -2453,16 +2383,6 @@ lean_dec(x_1); return x_11; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___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) { -_start: -{ -lean_object* x_11; -x_11 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_2); -lean_dec(x_1); -return x_11; -} -} LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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: { @@ -19997,6 +19917,62 @@ lean_dec(x_1); return x_3; } } +LEAN_EXPORT lean_object* l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___spec__1___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, 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; +x_17 = l_Lean_Elab_getDeclarationSelectionRef(x_1); +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(0, 7, 1); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +lean_ctor_set(x_19, 2, x_3); +lean_ctor_set(x_19, 3, x_4); +lean_ctor_set(x_19, 4, x_5); +lean_ctor_set(x_19, 5, x_6); +lean_ctor_set(x_19, 6, x_7); +lean_ctor_set_uint8(x_19, sizeof(void*)*7, x_2); +x_20 = lean_array_push(x_8, x_19); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_16); +return x_21; +} +} +static lean_object* _init_l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___spec__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("type of theorem '", 17); +return x_1; +} +} +static lean_object* _init_l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___spec__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___spec__1___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___spec__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("' is not a proposition", 22); +return x_1; +} +} +static lean_object* _init_l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___spec__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___spec__1___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} LEAN_EXPORT lean_object* l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___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) { _start: { @@ -20005,32 +19981,32 @@ x_14 = lean_unsigned_to_nat(0u); x_15 = lean_nat_dec_eq(x_5, x_14); if (x_15 == 0) { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_87; uint8_t x_88; +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_95; uint8_t x_96; x_16 = lean_unsigned_to_nat(1u); x_17 = lean_nat_sub(x_5, x_16); lean_dec(x_5); x_18 = lean_nat_sub(x_4, x_17); x_19 = lean_nat_sub(x_18, x_16); lean_dec(x_18); -x_87 = lean_array_get_size(x_2); -x_88 = lean_nat_dec_lt(x_19, x_87); -lean_dec(x_87); -if (x_88 == 0) +x_95 = lean_array_get_size(x_2); +x_96 = lean_nat_dec_lt(x_19, x_95); +lean_dec(x_95); +if (x_96 == 0) { -lean_object* x_89; lean_object* x_90; -x_89 = l_Lean_Elab_instInhabitedDefViewElabHeader; -x_90 = l___private_Init_Util_0__outOfBounds___rarg(x_89); -x_20 = x_90; -goto block_86; +lean_object* x_97; lean_object* x_98; +x_97 = l_Lean_Elab_instInhabitedDefViewElabHeader; +x_98 = l___private_Init_Util_0__outOfBounds___rarg(x_97); +x_20 = x_98; +goto block_94; } else { -lean_object* x_91; -x_91 = lean_array_fget(x_2, x_19); -x_20 = x_91; -goto block_86; +lean_object* x_99; +x_99 = lean_array_fget(x_2, x_19); +x_20 = x_99; +goto block_94; } -block_86: +block_94: { lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; x_21 = lean_ctor_get(x_20, 0); @@ -20056,105 +20032,171 @@ lean_inc(x_7); x_28 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint(x_27, x_7, x_8, x_9, x_10, x_11, x_12, x_13); if (lean_obj_tag(x_28) == 0) { -lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_85; uint8_t x_86; 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_array_get_size(x_3); -x_32 = lean_nat_dec_lt(x_19, x_31); -lean_dec(x_31); -if (x_32 == 0) +x_85 = lean_array_get_size(x_3); +x_86 = lean_nat_dec_lt(x_19, x_85); +lean_dec(x_85); +if (x_86 == 0) { -lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; uint8_t x_37; uint8_t x_38; lean_object* x_39; +lean_object* x_87; lean_object* x_88; lean_dec(x_19); -x_33 = l_Lean_instInhabitedExpr; -x_34 = l___private_Init_Util_0__outOfBounds___rarg(x_33); -x_35 = l_Lean_Elab_WF_TerminationHints_rememberExtraParams(x_25, x_29, x_34); +x_87 = l_Lean_instInhabitedExpr; +x_88 = l___private_Init_Util_0__outOfBounds___rarg(x_87); +x_31 = x_88; +goto block_84; +} +else +{ +lean_object* x_89; +x_89 = lean_array_fget(x_3, x_19); +lean_dec(x_19); +x_31 = x_89; +goto block_84; +} +block_84: +{ +lean_object* x_32; uint8_t x_33; uint8_t x_34; uint8_t x_35; lean_object* x_36; +x_32 = l_Lean_Elab_WF_TerminationHints_rememberExtraParams(x_25, x_29, x_31); lean_dec(x_25); -x_36 = 0; -x_37 = 1; -x_38 = 1; +x_33 = 0; +x_34 = 1; +x_35 = 1; lean_inc(x_1); -x_39 = l_Lean_Meta_mkLambdaFVars(x_1, x_34, x_36, x_37, x_38, x_9, x_10, x_11, x_12, x_30); +x_36 = l_Lean_Meta_mkLambdaFVars(x_1, x_31, x_33, x_34, x_35, x_9, x_10, x_11, x_12, x_30); +if (lean_obj_tag(x_36) == 0) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_36, 0); +lean_inc(x_37); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +lean_dec(x_36); +lean_inc(x_1); +x_39 = l_Lean_Meta_mkForallFVars(x_1, x_26, x_33, x_34, x_35, x_9, x_10, x_11, x_12, x_38); if (lean_obj_tag(x_39) == 0) { -lean_object* x_40; lean_object* x_41; lean_object* x_42; +lean_object* x_40; lean_object* x_41; uint8_t x_42; x_40 = lean_ctor_get(x_39, 0); lean_inc(x_40); x_41 = lean_ctor_get(x_39, 1); lean_inc(x_41); lean_dec(x_39); -lean_inc(x_1); -x_42 = l_Lean_Meta_mkForallFVars(x_1, x_26, x_36, x_37, x_38, x_9, x_10, x_11, x_12, x_41); -if (lean_obj_tag(x_42) == 0) +x_42 = l_Lean_Elab_DefKind_isTheorem(x_23); +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; -x_43 = lean_ctor_get(x_42, 0); -lean_inc(x_43); -x_44 = lean_ctor_get(x_42, 1); -lean_inc(x_44); -lean_dec(x_42); -x_45 = l_Lean_Elab_getDeclarationSelectionRef(x_21); -x_46 = lean_box(0); -x_47 = lean_alloc_ctor(0, 7, 1); -lean_ctor_set(x_47, 0, x_45); -lean_ctor_set(x_47, 1, x_46); -lean_ctor_set(x_47, 2, x_22); -lean_ctor_set(x_47, 3, x_24); -lean_ctor_set(x_47, 4, x_43); -lean_ctor_set(x_47, 5, x_40); -lean_ctor_set(x_47, 6, x_35); -lean_ctor_set_uint8(x_47, sizeof(void*)*7, x_23); -x_48 = lean_array_push(x_6, x_47); +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_43 = lean_box(0); +x_44 = l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___spec__1___lambda__1(x_21, x_23, x_22, x_24, x_40, x_37, x_32, x_6, x_43, x_7, x_8, x_9, x_10, x_11, x_12, x_41); +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_5 = x_17; -x_6 = x_48; -x_13 = x_44; +x_6 = x_45; +x_13 = x_46; goto _start; } else { -uint8_t x_50; -lean_dec(x_40); -lean_dec(x_35); -lean_dec(x_24); +lean_object* x_48; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_40); +x_48 = l_Lean_Meta_isProp(x_40, x_9, x_10, x_11, x_12, x_41); +if (lean_obj_tag(x_48) == 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_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; uint8_t x_62; +lean_dec(x_37); +lean_dec(x_32); lean_dec(x_22); -lean_dec(x_21); lean_dec(x_17); +lean_dec(x_6); +lean_dec(x_1); +x_51 = lean_ctor_get(x_48, 1); +lean_inc(x_51); +lean_dec(x_48); +x_52 = l_Lean_MessageData_ofName(x_24); +x_53 = l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___spec__1___closed__2; +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_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___spec__1___closed__4; +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 = l_Lean_indentExpr(x_40); +x_58 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_58, 0, x_56); +lean_ctor_set(x_58, 1, x_57); +x_59 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__2___closed__4; +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 = l_Lean_throwErrorAt___at_Lean_Elab_Term_elabMatch_elabMatchDefault___spec__7(x_21, x_60, x_7, x_8, x_9, x_10, x_11, x_12, x_51); 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_50 = !lean_is_exclusive(x_42); -if (x_50 == 0) +lean_dec(x_21); +x_62 = !lean_is_exclusive(x_61); +if (x_62 == 0) { -return x_42; +return x_61; } else { -lean_object* x_51; lean_object* x_52; lean_object* x_53; -x_51 = lean_ctor_get(x_42, 0); -x_52 = lean_ctor_get(x_42, 1); -lean_inc(x_52); -lean_inc(x_51); -lean_dec(x_42); -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; +lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_63 = lean_ctor_get(x_61, 0); +x_64 = lean_ctor_get(x_61, 1); +lean_inc(x_64); +lean_inc(x_63); +lean_dec(x_61); +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_48, 1); +lean_inc(x_66); +lean_dec(x_48); +x_67 = lean_box(0); +x_68 = l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___spec__1___lambda__1(x_21, x_23, x_22, x_24, x_40, x_37, x_32, x_6, x_67, x_7, x_8, x_9, x_10, x_11, x_12, x_66); +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_5 = x_17; +x_6 = x_69; +x_13 = x_70; +goto _start; } } else { -uint8_t x_54; -lean_dec(x_35); -lean_dec(x_26); +uint8_t x_72; +lean_dec(x_40); +lean_dec(x_37); +lean_dec(x_32); lean_dec(x_24); lean_dec(x_22); lean_dec(x_21); @@ -20167,78 +20209,32 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_1); -x_54 = !lean_is_exclusive(x_39); -if (x_54 == 0) +x_72 = !lean_is_exclusive(x_48); +if (x_72 == 0) { -return x_39; +return x_48; } else { -lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_55 = lean_ctor_get(x_39, 0); -x_56 = lean_ctor_get(x_39, 1); -lean_inc(x_56); -lean_inc(x_55); -lean_dec(x_39); -x_57 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_57, 0, x_55); -lean_ctor_set(x_57, 1, x_56); -return x_57; +lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_73 = lean_ctor_get(x_48, 0); +x_74 = lean_ctor_get(x_48, 1); +lean_inc(x_74); +lean_inc(x_73); +lean_dec(x_48); +x_75 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_75, 0, x_73); +lean_ctor_set(x_75, 1, x_74); +return x_75; } } } -else -{ -lean_object* x_58; lean_object* x_59; uint8_t x_60; uint8_t x_61; uint8_t x_62; lean_object* x_63; -x_58 = lean_array_fget(x_3, x_19); -lean_dec(x_19); -x_59 = l_Lean_Elab_WF_TerminationHints_rememberExtraParams(x_25, x_29, x_58); -lean_dec(x_25); -x_60 = 0; -x_61 = 1; -x_62 = 1; -lean_inc(x_1); -x_63 = l_Lean_Meta_mkLambdaFVars(x_1, x_58, x_60, x_61, x_62, x_9, x_10, x_11, x_12, x_30); -if (lean_obj_tag(x_63) == 0) -{ -lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_64 = lean_ctor_get(x_63, 0); -lean_inc(x_64); -x_65 = lean_ctor_get(x_63, 1); -lean_inc(x_65); -lean_dec(x_63); -lean_inc(x_1); -x_66 = l_Lean_Meta_mkForallFVars(x_1, x_26, x_60, x_61, x_62, x_9, x_10, x_11, x_12, x_65); -if (lean_obj_tag(x_66) == 0) -{ -lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; -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_Elab_getDeclarationSelectionRef(x_21); -x_70 = lean_box(0); -x_71 = lean_alloc_ctor(0, 7, 1); -lean_ctor_set(x_71, 0, x_69); -lean_ctor_set(x_71, 1, x_70); -lean_ctor_set(x_71, 2, x_22); -lean_ctor_set(x_71, 3, x_24); -lean_ctor_set(x_71, 4, x_67); -lean_ctor_set(x_71, 5, x_64); -lean_ctor_set(x_71, 6, x_59); -lean_ctor_set_uint8(x_71, sizeof(void*)*7, x_23); -x_72 = lean_array_push(x_6, x_71); -x_5 = x_17; -x_6 = x_72; -x_13 = x_68; -goto _start; } else { -uint8_t x_74; -lean_dec(x_64); -lean_dec(x_59); +uint8_t x_76; +lean_dec(x_37); +lean_dec(x_32); lean_dec(x_24); lean_dec(x_22); lean_dec(x_21); @@ -20251,30 +20247,30 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_1); -x_74 = !lean_is_exclusive(x_66); -if (x_74 == 0) +x_76 = !lean_is_exclusive(x_39); +if (x_76 == 0) { -return x_66; +return x_39; } else { -lean_object* x_75; lean_object* x_76; lean_object* x_77; -x_75 = lean_ctor_get(x_66, 0); -x_76 = lean_ctor_get(x_66, 1); -lean_inc(x_76); -lean_inc(x_75); -lean_dec(x_66); -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_77; lean_object* x_78; lean_object* x_79; +x_77 = lean_ctor_get(x_39, 0); +x_78 = lean_ctor_get(x_39, 1); +lean_inc(x_78); +lean_inc(x_77); +lean_dec(x_39); +x_79 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_79, 0, x_77); +lean_ctor_set(x_79, 1, x_78); +return x_79; } } } else { -uint8_t x_78; -lean_dec(x_59); +uint8_t x_80; +lean_dec(x_32); lean_dec(x_26); lean_dec(x_24); lean_dec(x_22); @@ -20288,30 +20284,30 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_1); -x_78 = !lean_is_exclusive(x_63); -if (x_78 == 0) +x_80 = !lean_is_exclusive(x_36); +if (x_80 == 0) { -return x_63; +return x_36; } else { -lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_79 = lean_ctor_get(x_63, 0); -x_80 = lean_ctor_get(x_63, 1); -lean_inc(x_80); -lean_inc(x_79); -lean_dec(x_63); -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; +lean_object* x_81; lean_object* x_82; lean_object* x_83; +x_81 = lean_ctor_get(x_36, 0); +x_82 = lean_ctor_get(x_36, 1); +lean_inc(x_82); +lean_inc(x_81); +lean_dec(x_36); +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_82; +uint8_t x_90; lean_dec(x_26); lean_dec(x_25); lean_dec(x_24); @@ -20327,30 +20323,30 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_1); -x_82 = !lean_is_exclusive(x_28); -if (x_82 == 0) +x_90 = !lean_is_exclusive(x_28); +if (x_90 == 0) { return x_28; } else { -lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_83 = lean_ctor_get(x_28, 0); -x_84 = lean_ctor_get(x_28, 1); -lean_inc(x_84); -lean_inc(x_83); +lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_91 = lean_ctor_get(x_28, 0); +x_92 = lean_ctor_get(x_28, 1); +lean_inc(x_92); +lean_inc(x_91); lean_dec(x_28); -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; +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_92; +lean_object* x_100; lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -20359,10 +20355,10 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); lean_dec(x_1); -x_92 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_92, 0, x_6); -lean_ctor_set(x_92, 1, x_13); -return x_92; +x_100 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_100, 0, x_6); +lean_ctor_set(x_100, 1, x_13); +return x_100; } } } @@ -20379,6 +20375,23 @@ lean_dec(x_3); return x_13; } } +LEAN_EXPORT lean_object* l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +uint8_t x_17; lean_object* x_18; +x_17 = lean_unbox(x_2); +lean_dec(x_2); +x_18 = l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___spec__1___lambda__1(x_1, x_17, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +return x_18; +} +} LEAN_EXPORT lean_object* l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___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) { _start: { @@ -20463,7 +20476,95 @@ return x_33; } } } -LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Elab_Term_MutualClosure_pushLetRecs___spec__1___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_EXPORT lean_object* l_List_foldlM___at_Lean_Elab_Term_MutualClosure_pushLetRecs___spec__1___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = lean_infer_type(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; uint8_t x_10; +x_9 = lean_ctor_get(x_7, 0); +x_10 = l_Lean_Expr_isProp(x_9); +lean_dec(x_9); +if (x_10 == 0) +{ +uint8_t x_11; lean_object* x_12; +x_11 = 0; +x_12 = lean_box(x_11); +lean_ctor_set(x_7, 0, x_12); +return x_7; +} +else +{ +uint8_t x_13; lean_object* x_14; +x_13 = 1; +x_14 = lean_box(x_13); +lean_ctor_set(x_7, 0, x_14); +return x_7; +} +} +else +{ +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); +lean_inc(x_16); +lean_inc(x_15); +lean_dec(x_7); +x_17 = l_Lean_Expr_isProp(x_15); +lean_dec(x_15); +if (x_17 == 0) +{ +uint8_t x_18; lean_object* x_19; lean_object* x_20; +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_16); +return x_20; +} +else +{ +uint8_t x_21; lean_object* x_22; lean_object* x_23; +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_16); +return x_23; +} +} +} +else +{ +uint8_t x_24; +x_24 = !lean_is_exclusive(x_7); +if (x_24 == 0) +{ +return x_7; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_7, 0); +x_26 = lean_ctor_get(x_7, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_7); +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_List_foldlM___at_Lean_Elab_Term_MutualClosure_pushLetRecs___spec__1___lambda__3(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; @@ -20595,30 +20696,32 @@ x_21 = l_Lean_Meta_Closure_mkLambda(x_15, x_19); x_22 = l_Lean_Elab_DefKind_isDefOrAbbrevOrOpaque(x_1); if (x_22 == 0) { -lean_object* x_23; lean_object* x_24; lean_object* x_25; +uint8_t x_23; +x_23 = l_Lean_Elab_DefKind_isTheorem(x_1); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_inc(x_2); -x_23 = l_List_foldlM___at_Lean_Elab_Term_MutualClosure_pushLetRecs___spec__1___lambda__1(x_2, x_12, x_14, x_20, x_21, x_3, x_1, x_5, x_6, x_7, x_8, x_9); +x_24 = l_List_foldlM___at_Lean_Elab_Term_MutualClosure_pushLetRecs___spec__1___lambda__1(x_2, x_12, x_14, x_20, x_21, x_3, x_1, x_5, x_6, x_7, x_8, x_9); lean_dec(x_12); -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); +x_25 = lean_ctor_get(x_24, 0); lean_inc(x_25); -lean_dec(x_23); -x_3 = x_24; +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +lean_dec(x_24); +x_3 = x_25; x_4 = x_13; -x_9 = x_25; +x_9 = x_26; goto _start; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_box(x_1); -x_28 = lean_alloc_closure((void*)(l_List_foldlM___at_Lean_Elab_Term_MutualClosure_pushLetRecs___spec__1___lambda__2___boxed), 7, 2); +lean_object* x_28; lean_object* x_29; +x_28 = lean_alloc_closure((void*)(l_List_foldlM___at_Lean_Elab_Term_MutualClosure_pushLetRecs___spec__1___lambda__2), 6, 1); lean_closure_set(x_28, 0, x_18); -lean_closure_set(x_28, 1, x_27); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); @@ -20682,6 +20785,76 @@ return x_40; } } } +else +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_box(x_1); +x_42 = lean_alloc_closure((void*)(l_List_foldlM___at_Lean_Elab_Term_MutualClosure_pushLetRecs___spec__1___lambda__3___boxed), 7, 2); +lean_closure_set(x_42, 0, x_18); +lean_closure_set(x_42, 1, x_41); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_43 = l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___rarg(x_16, x_17, x_42, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_43) == 0) +{ +lean_object* x_44; lean_object* x_45; uint8_t x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_44 = lean_ctor_get(x_43, 0); +lean_inc(x_44); +x_45 = lean_ctor_get(x_43, 1); +lean_inc(x_45); +lean_dec(x_43); +x_46 = lean_unbox(x_44); +lean_dec(x_44); +lean_inc(x_2); +x_47 = l_List_foldlM___at_Lean_Elab_Term_MutualClosure_pushLetRecs___spec__1___lambda__1(x_2, x_12, x_14, x_20, x_21, x_3, x_46, x_5, x_6, x_7, x_8, x_45); +lean_dec(x_12); +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_3 = x_48; +x_4 = x_13; +x_9 = x_49; +goto _start; +} +else +{ +uint8_t x_51; +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +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_51 = !lean_is_exclusive(x_43); +if (x_51 == 0) +{ +return x_43; +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_43, 0); +x_53 = lean_ctor_get(x_43, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_43); +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_EXPORT lean_object* l_Lean_Elab_Term_MutualClosure_pushLetRecs(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) { @@ -20707,13 +20880,13 @@ lean_dec(x_2); return x_14; } } -LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Elab_Term_MutualClosure_pushLetRecs___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_EXPORT lean_object* l_List_foldlM___at_Lean_Elab_Term_MutualClosure_pushLetRecs___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) { _start: { uint8_t x_8; lean_object* x_9; x_8 = lean_unbox(x_2); lean_dec(x_2); -x_9 = l_List_foldlM___at_Lean_Elab_Term_MutualClosure_pushLetRecs___spec__1___lambda__2(x_1, x_8, x_3, x_4, x_5, x_6, x_7); +x_9 = l_List_foldlM___at_Lean_Elab_Term_MutualClosure_pushLetRecs___spec__1___lambda__3(x_1, x_8, x_3, x_4, x_5, x_6, x_7); return x_9; } } @@ -34000,10 +34173,6 @@ l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__6___closed__1 lean_mark_persistent(l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__6___closed__1); l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__6___closed__2 = _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__6___closed__2(); lean_mark_persistent(l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__6___closed__2); -l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__7___closed__1 = _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__7___closed__1(); -lean_mark_persistent(l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__7___closed__1); -l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__7___closed__2 = _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__7___closed__2(); -lean_mark_persistent(l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__7___closed__2); l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___closed__1 = _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___closed__1(); lean_mark_persistent(l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___closed__1); l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___closed__2 = _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___closed__2(); @@ -34327,6 +34496,14 @@ l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkClosureForAux_ lean_mark_persistent(l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkClosureForAux___closed__6); l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosureFor___lambda__2___closed__1 = _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosureFor___lambda__2___closed__1(); lean_mark_persistent(l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosureFor___lambda__2___closed__1); +l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___spec__1___closed__1 = _init_l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___spec__1___closed__1(); +lean_mark_persistent(l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___spec__1___closed__1); +l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___spec__1___closed__2 = _init_l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___spec__1___closed__2(); +lean_mark_persistent(l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___spec__1___closed__2); +l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___spec__1___closed__3 = _init_l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___spec__1___closed__3(); +lean_mark_persistent(l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___spec__1___closed__3); +l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___spec__1___closed__4 = _init_l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___spec__1___closed__4(); +lean_mark_persistent(l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___spec__1___closed__4); l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_levelMVarToParamHeaders_process___spec__1___closed__1 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_levelMVarToParamHeaders_process___spec__1___closed__1(); lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_levelMVarToParamHeaders_process___spec__1___closed__1); l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__1 = _init_l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__1(); diff --git a/stage0/stdlib/Lean/Elab/Tactic/Omega/Core.c b/stage0/stdlib/Lean/Elab/Tactic/Omega/Core.c index c84919d34a23..efc231d30d99 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Omega/Core.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Omega/Core.c @@ -43,7 +43,7 @@ static lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omeg LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Elab_Tactic_Omega_Problem_fourierMotzkinSelect___spec__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_solveEasyEquality(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_Justification_combineProof___closed__3; -lean_object* l___private_Init_Omega_Constraint_0__Lean_Omega_decEqConstraint____x40_Init_Omega_Constraint___hyg_142_(lean_object*, lean_object*); +uint8_t l___private_Init_Omega_Constraint_0__Lean_Omega_decEqConstraint____x40_Init_Omega_Constraint___hyg_142_(lean_object*, lean_object*); static lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1___closed__4; static lean_object* l_Lean_Elab_Tactic_Omega_instToExprConstraint___closed__4; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Tactic_Omega_Problem_fourierMotzkinData___spec__2___lambda__1___boxed(lean_object*, lean_object*); @@ -75,7 +75,6 @@ static lean_object* l_Lean_Elab_Tactic_Omega_Justification_toString___closed__10 LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_addConstraint(lean_object*, lean_object*); size_t lean_hashset_mk_idx(lean_object*, uint64_t); LEAN_EXPORT lean_object* l_List_replace___at_Lean_Elab_Tactic_Omega_Problem_insertConstraint___spec__15___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Omega_Problem_addConstraint___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_addEqualities(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___closed__3; static lean_object* l_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___closed__7; @@ -118,7 +117,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_selectEquality(lean_ob lean_object* l_ReaderT_instMonadReaderT___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Elab_Tactic_Omega_Problem_solveEasyEquality___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_instToExprConstraint___lambda__1___closed__8; -uint8_t l_List_hasDecEq___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___lambda__1___closed__13; static lean_object* l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___closed__1; lean_object* l_Lean_Omega_Constraint_exact(lean_object*); @@ -344,6 +342,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_runOmega(lean_object*, static uint8_t l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Omega_Problem_fourierMotzkinSelect___spec__4___closed__2; static lean_object* l_Lean_Elab_Tactic_Omega_Problem_addEquality___closed__2; static lean_object* l_Lean_Elab_Tactic_Omega_Justification_toString___closed__4; +uint8_t l_List_hasDecEq___at___private_Init_Omega_LinearCombo_0__Lean_Omega_decEqLinearCombo____x40_Init_Omega_LinearCombo___hyg_27____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_instToExprConstraint___lambda__1(lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_Justification_bmodProof___closed__9; static lean_object* l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___lambda__1___closed__10; @@ -493,7 +492,6 @@ LEAN_EXPORT uint8_t l_Lean_Elab_Tactic_Omega_Problem_FourierMotzkinData_exact(le LEAN_EXPORT uint64_t l_List_foldl___at_Lean_Elab_Tactic_Omega_Problem_insertConstraint___spec__2(uint64_t, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_Problem_proveFalse___closed__7; -lean_object* l_Int_decEq___boxed(lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); static lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Core___hyg_5____closed__19; lean_object* lean_nat_add(lean_object*, lean_object*); @@ -8632,14 +8630,6 @@ lean_dec(x_2); return x_9; } } -static lean_object* _init_l_Lean_Elab_Tactic_Omega_Problem_addConstraint___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Int_decEq___boxed), 2, 0); -return x_1; -} -} LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_addConstraint(lean_object* x_1, lean_object* x_2) { _start: { @@ -8708,14 +8698,13 @@ lean_dec(x_8); 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; uint8_t x_19; +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; x_15 = lean_ctor_get(x_13, 0); x_16 = lean_ctor_get(x_13, 1); x_17 = lean_ctor_get(x_13, 2); -x_18 = l_Lean_Elab_Tactic_Omega_Problem_addConstraint___closed__1; -lean_inc(x_4); -x_19 = l_List_hasDecEq___rarg(x_18, x_4, x_15); -if (x_19 == 0) +x_18 = l_List_hasDecEq___at___private_Init_Omega_LinearCombo_0__Lean_Omega_decEqLinearCombo____x40_Init_Omega_LinearCombo___hyg_27____spec__1(x_4, x_15); +lean_dec(x_15); +if (x_18 == 0) { lean_free_object(x_13); lean_dec(x_17); @@ -8727,55 +8716,47 @@ return x_1; } else { -lean_object* x_20; lean_object* x_21; uint8_t x_22; +lean_object* x_19; uint8_t x_20; lean_inc(x_16); lean_inc(x_5); -x_20 = l_Lean_Omega_Constraint_combine(x_5, x_16); -lean_inc(x_16); -lean_inc(x_20); -x_21 = l___private_Init_Omega_Constraint_0__Lean_Omega_decEqConstraint____x40_Init_Omega_Constraint___hyg_142_(x_20, x_16); -x_22 = lean_unbox(x_21); -lean_dec(x_21); -if (x_22 == 0) +x_19 = l_Lean_Omega_Constraint_combine(x_5, x_16); +x_20 = l___private_Init_Omega_Constraint_0__Lean_Omega_decEqConstraint____x40_Init_Omega_Constraint___hyg_142_(x_19, x_16); +if (x_20 == 0) { -lean_object* x_23; uint8_t x_24; -lean_inc(x_5); -lean_inc(x_20); -x_23 = l___private_Init_Omega_Constraint_0__Lean_Omega_decEqConstraint____x40_Init_Omega_Constraint___hyg_142_(x_20, x_5); -x_24 = lean_unbox(x_23); -lean_dec(x_23); -if (x_24 == 0) +uint8_t x_21; +x_21 = l___private_Init_Omega_Constraint_0__Lean_Omega_decEqConstraint____x40_Init_Omega_Constraint___hyg_142_(x_19, x_5); +if (x_21 == 0) { -lean_object* x_25; lean_object* x_26; +lean_object* x_22; lean_object* x_23; lean_inc(x_4); -x_25 = lean_alloc_ctor(2, 5, 0); -lean_ctor_set(x_25, 0, x_5); -lean_ctor_set(x_25, 1, x_16); -lean_ctor_set(x_25, 2, x_4); -lean_ctor_set(x_25, 3, x_6); -lean_ctor_set(x_25, 4, x_17); -lean_ctor_set(x_13, 2, x_25); -lean_ctor_set(x_13, 1, x_20); +x_22 = lean_alloc_ctor(2, 5, 0); +lean_ctor_set(x_22, 0, x_5); +lean_ctor_set(x_22, 1, x_16); +lean_ctor_set(x_22, 2, x_4); +lean_ctor_set(x_22, 3, x_6); +lean_ctor_set(x_22, 4, x_17); +lean_ctor_set(x_13, 2, x_22); +lean_ctor_set(x_13, 1, x_19); lean_ctor_set(x_13, 0, x_4); -x_26 = l_Lean_Elab_Tactic_Omega_Problem_insertConstraint(x_1, x_13); -return x_26; +x_23 = l_Lean_Elab_Tactic_Omega_Problem_insertConstraint(x_1, x_13); +return x_23; } else { -lean_object* x_27; -lean_dec(x_20); +lean_object* x_24; +lean_dec(x_19); lean_dec(x_17); lean_dec(x_16); lean_ctor_set(x_13, 2, x_6); lean_ctor_set(x_13, 1, x_5); lean_ctor_set(x_13, 0, x_4); -x_27 = l_Lean_Elab_Tactic_Omega_Problem_insertConstraint(x_1, x_13); -return x_27; +x_24 = l_Lean_Elab_Tactic_Omega_Problem_insertConstraint(x_1, x_13); +return x_24; } } else { -lean_dec(x_20); +lean_dec(x_19); lean_free_object(x_13); lean_dec(x_17); lean_dec(x_16); @@ -8788,21 +8769,20 @@ return x_1; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; -x_28 = lean_ctor_get(x_13, 0); -x_29 = lean_ctor_get(x_13, 1); -x_30 = lean_ctor_get(x_13, 2); -lean_inc(x_30); -lean_inc(x_29); -lean_inc(x_28); +lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_25 = lean_ctor_get(x_13, 0); +x_26 = lean_ctor_get(x_13, 1); +x_27 = lean_ctor_get(x_13, 2); +lean_inc(x_27); +lean_inc(x_26); +lean_inc(x_25); lean_dec(x_13); -x_31 = l_Lean_Elab_Tactic_Omega_Problem_addConstraint___closed__1; -lean_inc(x_4); -x_32 = l_List_hasDecEq___rarg(x_31, x_4, x_28); -if (x_32 == 0) +x_28 = l_List_hasDecEq___at___private_Init_Omega_LinearCombo_0__Lean_Omega_decEqLinearCombo____x40_Init_Omega_LinearCombo___hyg_27____spec__1(x_4, x_25); +lean_dec(x_25); +if (x_28 == 0) { -lean_dec(x_30); -lean_dec(x_29); +lean_dec(x_27); +lean_dec(x_26); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); @@ -8810,59 +8790,51 @@ return x_1; } else { -lean_object* x_33; lean_object* x_34; uint8_t x_35; -lean_inc(x_29); +lean_object* x_29; uint8_t x_30; +lean_inc(x_26); lean_inc(x_5); -x_33 = l_Lean_Omega_Constraint_combine(x_5, x_29); -lean_inc(x_29); -lean_inc(x_33); -x_34 = l___private_Init_Omega_Constraint_0__Lean_Omega_decEqConstraint____x40_Init_Omega_Constraint___hyg_142_(x_33, x_29); -x_35 = lean_unbox(x_34); -lean_dec(x_34); -if (x_35 == 0) +x_29 = l_Lean_Omega_Constraint_combine(x_5, x_26); +x_30 = l___private_Init_Omega_Constraint_0__Lean_Omega_decEqConstraint____x40_Init_Omega_Constraint___hyg_142_(x_29, x_26); +if (x_30 == 0) { -lean_object* x_36; uint8_t x_37; -lean_inc(x_5); -lean_inc(x_33); -x_36 = l___private_Init_Omega_Constraint_0__Lean_Omega_decEqConstraint____x40_Init_Omega_Constraint___hyg_142_(x_33, x_5); -x_37 = lean_unbox(x_36); -lean_dec(x_36); -if (x_37 == 0) +uint8_t x_31; +x_31 = l___private_Init_Omega_Constraint_0__Lean_Omega_decEqConstraint____x40_Init_Omega_Constraint___hyg_142_(x_29, x_5); +if (x_31 == 0) { -lean_object* x_38; lean_object* x_39; lean_object* x_40; +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_inc(x_4); -x_38 = lean_alloc_ctor(2, 5, 0); -lean_ctor_set(x_38, 0, x_5); -lean_ctor_set(x_38, 1, x_29); -lean_ctor_set(x_38, 2, x_4); -lean_ctor_set(x_38, 3, x_6); -lean_ctor_set(x_38, 4, x_30); -x_39 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_39, 0, x_4); -lean_ctor_set(x_39, 1, x_33); -lean_ctor_set(x_39, 2, x_38); -x_40 = l_Lean_Elab_Tactic_Omega_Problem_insertConstraint(x_1, x_39); -return x_40; +x_32 = lean_alloc_ctor(2, 5, 0); +lean_ctor_set(x_32, 0, x_5); +lean_ctor_set(x_32, 1, x_26); +lean_ctor_set(x_32, 2, x_4); +lean_ctor_set(x_32, 3, x_6); +lean_ctor_set(x_32, 4, x_27); +x_33 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_33, 0, x_4); +lean_ctor_set(x_33, 1, x_29); +lean_ctor_set(x_33, 2, x_32); +x_34 = l_Lean_Elab_Tactic_Omega_Problem_insertConstraint(x_1, x_33); +return x_34; } else { -lean_object* x_41; lean_object* x_42; -lean_dec(x_33); -lean_dec(x_30); +lean_object* x_35; lean_object* x_36; lean_dec(x_29); -x_41 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_41, 0, x_4); -lean_ctor_set(x_41, 1, x_5); -lean_ctor_set(x_41, 2, x_6); -x_42 = l_Lean_Elab_Tactic_Omega_Problem_insertConstraint(x_1, x_41); -return x_42; +lean_dec(x_27); +lean_dec(x_26); +x_35 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_35, 0, x_4); +lean_ctor_set(x_35, 1, x_5); +lean_ctor_set(x_35, 2, x_6); +x_36 = l_Lean_Elab_Tactic_Omega_Problem_insertConstraint(x_1, x_35); +return x_36; } } else { -lean_dec(x_33); -lean_dec(x_30); lean_dec(x_29); +lean_dec(x_27); +lean_dec(x_26); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); @@ -16451,8 +16423,6 @@ l_Lean_Elab_Tactic_Omega_Problem_proveFalse___closed__10 = _init_l_Lean_Elab_Tac lean_mark_persistent(l_Lean_Elab_Tactic_Omega_Problem_proveFalse___closed__10); l_Lean_Elab_Tactic_Omega_Problem_proveFalse___closed__11 = _init_l_Lean_Elab_Tactic_Omega_Problem_proveFalse___closed__11(); lean_mark_persistent(l_Lean_Elab_Tactic_Omega_Problem_proveFalse___closed__11); -l_Lean_Elab_Tactic_Omega_Problem_addConstraint___closed__1 = _init_l_Lean_Elab_Tactic_Omega_Problem_addConstraint___closed__1(); -lean_mark_persistent(l_Lean_Elab_Tactic_Omega_Problem_addConstraint___closed__1); l_List_foldr___at_Lean_Elab_Tactic_Omega_Problem_replayEliminations___spec__1___closed__1 = _init_l_List_foldr___at_Lean_Elab_Tactic_Omega_Problem_replayEliminations___spec__1___closed__1(); lean_mark_persistent(l_List_foldr___at_Lean_Elab_Tactic_Omega_Problem_replayEliminations___spec__1___closed__1); l_List_foldr___at_Lean_Elab_Tactic_Omega_Problem_replayEliminations___spec__1___closed__2 = _init_l_List_foldr___at_Lean_Elab_Tactic_Omega_Problem_replayEliminations___spec__1___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/Tactic/Omega/MinNatAbs.c b/stage0/stdlib/Lean/Elab/Tactic/Omega/MinNatAbs.c index 6c36350a57dd..05de9b577eba 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Omega/MinNatAbs.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Omega/MinNatAbs.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Elab.Tactic.Omega.MinNatAbs -// Imports: Init.BinderPredicates Init.Data.Option.Lemmas Init.Data.Nat.Bitwise.Lemmas +// Imports: Init.BinderPredicates Init.Data.Int.Order Init.Data.List.Lemmas Init.Data.Nat.MinMax Init.Data.Option.Lemmas #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -347,8 +347,10 @@ return x_2; } } lean_object* initialize_Init_BinderPredicates(uint8_t builtin, lean_object*); +lean_object* initialize_Init_Data_Int_Order(uint8_t builtin, lean_object*); +lean_object* initialize_Init_Data_List_Lemmas(uint8_t builtin, lean_object*); +lean_object* initialize_Init_Data_Nat_MinMax(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_Option_Lemmas(uint8_t builtin, lean_object*); -lean_object* initialize_Init_Data_Nat_Bitwise_Lemmas(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Elab_Tactic_Omega_MinNatAbs(uint8_t builtin, lean_object* w) { lean_object * res; @@ -357,10 +359,16 @@ _G_initialized = true; res = initialize_Init_BinderPredicates(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -res = initialize_Init_Data_Option_Lemmas(builtin, lean_io_mk_world()); +res = initialize_Init_Data_Int_Order(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Init_Data_List_Lemmas(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -res = initialize_Init_Data_Nat_Bitwise_Lemmas(builtin, lean_io_mk_world()); +res = initialize_Init_Data_Nat_MinMax(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Init_Data_Option_Lemmas(builtin, 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/Omega/OmegaM.c b/stage0/stdlib/Lean/Elab/Tactic/Omega/OmegaM.c index 6803bb6b8757..a1d5901307ad 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Omega/OmegaM.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Omega/OmegaM.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Elab.Tactic.Omega.OmegaM -// Imports: Init.Omega.LinearCombo Init.Omega.Int Init.Omega.Logic Init.Data.BitVec Lean.Meta.AppBuilder +// Imports: Init.Omega.LinearCombo Init.Omega.Int Init.Omega.Logic Init.Data.BitVec.Basic Lean.Meta.AppBuilder #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -24,7 +24,6 @@ static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__75; lean_object* l_Lean_mkNatLit(lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__27; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_atomsList___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__110; static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__79; LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_Elab_Tactic_Omega_lookup___spec__2(lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); @@ -3626,7 +3625,7 @@ static lean_object* _init_l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__94() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("toNat_lt", 8); +x_1 = lean_mk_string_from_bytes("isLt", 4); return x_1; } } @@ -3661,32 +3660,24 @@ return x_1; static lean_object* _init_l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__98() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("isLt", 4); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__99() { -_start: -{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__91; -x_2 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__98; +x_2 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__94; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__100() { +static lean_object* _init_l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__99() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__99; +x_2 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__98; x_3 = l_Lean_Expr_const___override(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__101() { +static lean_object* _init_l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__100() { _start: { lean_object* x_1; @@ -3694,7 +3685,7 @@ x_1 = lean_mk_string_from_bytes("natAbs", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__102() { +static lean_object* _init_l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__101() { _start: { lean_object* x_1; @@ -3702,27 +3693,27 @@ x_1 = lean_mk_string_from_bytes("le_natAbs", 9); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__103() { +static lean_object* _init_l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__102() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Tactic_Omega_atomsList___rarg___closed__1; -x_2 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__102; +x_2 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__101; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__104() { +static lean_object* _init_l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__103() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__103; +x_2 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__102; x_3 = l_Lean_Expr_const___override(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__105() { +static lean_object* _init_l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__104() { _start: { lean_object* x_1; @@ -3730,29 +3721,29 @@ x_1 = lean_mk_string_from_bytes("neg_le_natAbs", 13); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__106() { +static lean_object* _init_l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__105() { _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_Tactic_Omega_atomsCoeffs___rarg___closed__1; x_2 = l_Lean_Elab_Tactic_Omega_atomsCoeffs___rarg___closed__2; x_3 = l_Lean_Elab_Tactic_Omega_atomsList___rarg___closed__1; -x_4 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__105; +x_4 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__104; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__107() { +static lean_object* _init_l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__106() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__106; +x_2 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__105; x_3 = l_Lean_Expr_const___override(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__108() { +static lean_object* _init_l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__107() { _start: { lean_object* x_1; @@ -3760,24 +3751,24 @@ x_1 = lean_mk_string_from_bytes("ofNat_sub_dichotomy", 19); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__109() { +static lean_object* _init_l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__108() { _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_Tactic_Omega_atomsCoeffs___rarg___closed__1; x_2 = l_Lean_Elab_Tactic_Omega_atomsCoeffs___rarg___closed__2; x_3 = l_Lean_Elab_Tactic_Omega_atomsList___rarg___closed__1; -x_4 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__108; +x_4 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__107; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__110() { +static lean_object* _init_l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__109() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__109; +x_2 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__108; x_3 = l_Lean_Expr_const___override(x_2, x_1); return x_3; } @@ -5420,7 +5411,7 @@ x_436 = lean_array_fget(x_408, x_387); x_437 = lean_unsigned_to_nat(1u); x_438 = lean_array_fget(x_408, x_437); lean_dec(x_408); -x_439 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__100; +x_439 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__99; x_440 = l_Lean_mkAppB(x_439, x_436, x_438); x_441 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_402, x_440); x_442 = lean_alloc_ctor(0, 2, 0); @@ -5435,7 +5426,7 @@ else { lean_object* x_443; uint8_t x_444; lean_dec(x_410); -x_443 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__101; +x_443 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__100; x_444 = lean_string_dec_eq(x_409, x_443); lean_dec(x_409); if (x_444 == 0) @@ -5468,10 +5459,10 @@ else 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; x_450 = lean_array_fget(x_408, x_387); lean_dec(x_408); -x_451 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__104; +x_451 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__103; lean_inc(x_450); x_452 = l_Lean_Expr_app___override(x_451, x_450); -x_453 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__107; +x_453 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__106; x_454 = l_Lean_Expr_app___override(x_453, x_450); x_455 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_402, x_452); x_456 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_455, x_454); @@ -5659,7 +5650,7 @@ x_494 = lean_array_fget(x_464, x_387); x_495 = lean_unsigned_to_nat(1u); x_496 = lean_array_fget(x_464, x_495); lean_dec(x_464); -x_497 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__100; +x_497 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__99; x_498 = l_Lean_mkAppB(x_497, x_494, x_496); x_499 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_402, x_498); x_500 = lean_alloc_ctor(0, 2, 0); @@ -5674,7 +5665,7 @@ else { lean_object* x_501; uint8_t x_502; lean_dec(x_466); -x_501 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__101; +x_501 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__100; x_502 = lean_string_dec_eq(x_465, x_501); lean_dec(x_465); if (x_502 == 0) @@ -5707,10 +5698,10 @@ else 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; x_508 = lean_array_fget(x_464, x_387); lean_dec(x_464); -x_509 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__104; +x_509 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__103; lean_inc(x_508); x_510 = l_Lean_Expr_app___override(x_509, x_508); -x_511 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__107; +x_511 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__106; x_512 = l_Lean_Expr_app___override(x_511, x_508); x_513 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_402, x_510); x_514 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_513, x_512); @@ -5762,7 +5753,7 @@ x_524 = lean_array_fget(x_464, x_523); x_525 = lean_unsigned_to_nat(5u); x_526 = lean_array_fget(x_464, x_525); lean_dec(x_464); -x_527 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__110; +x_527 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__109; x_528 = l_Lean_mkAppB(x_527, x_524, x_526); x_529 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_402, x_528); x_530 = lean_alloc_ctor(0, 2, 0); @@ -7290,7 +7281,7 @@ return x_12; lean_object* initialize_Init_Omega_LinearCombo(uint8_t builtin, lean_object*); lean_object* initialize_Init_Omega_Int(uint8_t builtin, lean_object*); lean_object* initialize_Init_Omega_Logic(uint8_t builtin, lean_object*); -lean_object* initialize_Init_Data_BitVec(uint8_t builtin, lean_object*); +lean_object* initialize_Init_Data_BitVec_Basic(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_AppBuilder(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Elab_Tactic_Omega_OmegaM(uint8_t builtin, lean_object* w) { @@ -7306,7 +7297,7 @@ lean_dec_ref(res); res = initialize_Init_Omega_Logic(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -res = initialize_Init_Data_BitVec(builtin, lean_io_mk_world()); +res = initialize_Init_Data_BitVec_Basic(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); res = initialize_Lean_Meta_AppBuilder(builtin, lean_io_mk_world()); @@ -7599,8 +7590,6 @@ l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__108 = _init_l_Lean_Elab_Tactic_Om lean_mark_persistent(l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__108); l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__109 = _init_l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__109(); lean_mark_persistent(l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__109); -l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__110 = _init_l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__110(); -lean_mark_persistent(l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__110); l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3___closed__1 = _init_l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3___closed__1(); lean_mark_persistent(l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3___closed__1); l_Lean_Elab_Tactic_Omega_lookup___lambda__2___closed__1 = _init_l_Lean_Elab_Tactic_Omega_lookup___lambda__2___closed__1(); diff --git a/stage0/stdlib/Lean/Elab/Tactic/SimpTrace.c b/stage0/stdlib/Lean/Elab/Tactic/SimpTrace.c index e52b580efa27..ef93edfc45f3 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/SimpTrace.c +++ b/stage0/stdlib/Lean/Elab/Tactic/SimpTrace.c @@ -48,7 +48,7 @@ lean_object* l_Lean_addBuiltinDeclarationRanges(lean_object*, lean_object*, lean static lean_object* l_Lean_Elab_Tactic_evalSimpTrace___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalSimpTrace___closed__3; lean_object* l_Lean_Meta_Tactic_TryThis_addSuggestion(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_dsimpLocation_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_dsimpLocation_x27(lean_object*, 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_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*); lean_object* l_Lean_Elab_Tactic_Simp_DischargeWrapper_with___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -60,13 +60,13 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalSimpAllTrace___lambda__2(lean_ob static lean_object* l_Lean_Elab_Tactic_evalSimpTrace___lambda__4___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalSimpAllTrace___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___regBuiltin_Lean_Elab_Tactic_evalSimpAllTrace___closed__2; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_dsimpLocation_x27___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_dsimpLocation_x27___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_evalSimpAllTrace___lambda__5___closed__1; lean_object* l_Lean_MVarId_getNondepPropHyps(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_evalSimpTrace___lambda__7___closed__1; static lean_object* l_Lean_Elab_Tactic_evalSimpTrace___lambda__2___closed__2; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_dsimpLocation_x27_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_dsimpLocation_x27_go(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_dsimpLocation_x27_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_dsimpLocation_x27_go(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_evalSimpAllTrace___lambda__3___closed__4; lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); lean_object* l_Lean_Syntax_node6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -124,7 +124,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalSimpTrace___lambda__5(lean_objec LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalSimpAllTrace___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* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalExact___spec__1___rarg(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_evalDSimpTrace(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_dsimpLocation_x27___lambda__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_dsimpLocation_x27___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*); uint8_t l_Lean_Syntax_isNone(lean_object*); static lean_object* l_Lean_Elab_Tactic_evalSimpAllTrace___lambda__2___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalSimpAllTrace___closed__3; @@ -153,7 +153,7 @@ lean_object* l_Lean_Elab_Tactic_mkSimpOnly(lean_object*, lean_object*, lean_obje LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_dsimpLocation_x27_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*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_evalSimpTrace_declRange(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalDSimpTrace___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_Lean_Elab_Tactic_dsimpLocation_x27___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_Tactic_dsimpLocation_x27___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_Lean_Elab_Tactic_evalSimpAllTrace___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalSimpAllTrace_declRange___closed__7; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalDSimpTrace___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*); @@ -2955,10 +2955,11 @@ lean_ctor_set(x_12, 1, x_11); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_dsimpLocation_x27_go(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_EXPORT lean_object* l_Lean_Elab_Tactic_dsimpLocation_x27_go(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_13; +lean_object* x_14; +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); @@ -2966,23 +2967,21 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_4); -x_13 = l_Lean_Elab_Tactic_getMainGoal(x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_13) == 0) +x_14 = l_Lean_Elab_Tactic_getMainGoal(x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, 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; -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_15 = lean_ctor_get(x_14, 0); lean_inc(x_15); -lean_dec(x_13); -x_16 = l_Lean_Elab_Tactic_evalSimpTrace___lambda__1___closed__1; +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); x_17 = l_Lean_Elab_Tactic_evalSimpAllTrace___lambda__2___closed__1; +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); -lean_inc(x_8); -x_18 = l_Lean_Meta_dsimpGoal(x_14, x_1, x_16, x_3, x_2, x_17, x_8, x_9, x_10, x_11, x_15); +x_18 = l_Lean_Meta_dsimpGoal(x_15, x_1, x_2, x_4, x_3, x_17, x_9, x_10, x_11, x_12, x_16); if (lean_obj_tag(x_18) == 0) { lean_object* x_19; lean_object* x_20; @@ -3000,7 +2999,7 @@ x_22 = lean_ctor_get(x_19, 1); lean_inc(x_22); lean_dec(x_19); x_23 = lean_box(0); -x_24 = l_Lean_Elab_Tactic_replaceMainGoal(x_23, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_21); +x_24 = l_Lean_Elab_Tactic_replaceMainGoal(x_23, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_21); if (lean_obj_tag(x_24) == 0) { uint8_t x_25; @@ -3065,7 +3064,7 @@ x_36 = lean_box(0); x_37 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_37, 0, x_35); lean_ctor_set(x_37, 1, x_36); -x_38 = l_Lean_Elab_Tactic_replaceMainGoal(x_37, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_33); +x_38 = l_Lean_Elab_Tactic_replaceMainGoal(x_37, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_33); if (lean_obj_tag(x_38) == 0) { uint8_t x_39; @@ -3118,6 +3117,7 @@ return x_46; else { uint8_t x_47; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -3125,7 +3125,6 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); x_47 = !lean_is_exclusive(x_18); if (x_47 == 0) { @@ -3149,6 +3148,7 @@ return x_50; else { uint8_t x_51; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -3156,22 +3156,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_2); lean_dec(x_1); -x_51 = !lean_is_exclusive(x_13); +x_51 = !lean_is_exclusive(x_14); if (x_51 == 0) { -return x_13; +return x_14; } else { lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_ctor_get(x_13, 0); -x_53 = lean_ctor_get(x_13, 1); +x_52 = lean_ctor_get(x_14, 0); +x_53 = lean_ctor_get(x_14, 1); lean_inc(x_53); lean_inc(x_52); -lean_dec(x_13); +lean_dec(x_14); x_54 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_54, 0, x_52); lean_ctor_set(x_54, 1, x_53); @@ -3197,20 +3197,21 @@ lean_dec(x_2); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_dsimpLocation_x27_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, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_dsimpLocation_x27_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, lean_object* x_12, lean_object* x_13) { _start: { -uint8_t x_13; lean_object* x_14; -x_13 = lean_unbox(x_3); -lean_dec(x_3); -x_14 = l_Lean_Elab_Tactic_dsimpLocation_x27_go(x_1, x_2, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_14; +uint8_t x_14; lean_object* x_15; +x_14 = lean_unbox(x_4); +lean_dec(x_4); +x_15 = l_Lean_Elab_Tactic_dsimpLocation_x27_go(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_Lean_Elab_Tactic_dsimpLocation_x27___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_Elab_Tactic_dsimpLocation_x27___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_11; +lean_object* x_12; +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); @@ -3218,36 +3219,36 @@ lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -lean_inc(x_2); -x_11 = l_Lean_Elab_Tactic_getMainGoal(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) +x_12 = l_Lean_Elab_Tactic_getMainGoal(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_12) == 0) { -lean_object* x_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_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_12, 0); lean_inc(x_13); -lean_dec(x_11); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -lean_inc(x_6); -x_14 = l_Lean_MVarId_getNondepPropHyps(x_12, x_6, x_7, x_8, x_9, x_13); -if (lean_obj_tag(x_14) == 0) +x_15 = l_Lean_MVarId_getNondepPropHyps(x_13, x_7, x_8, x_9, x_10, x_14); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_14, 1); +lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; +x_16 = lean_ctor_get(x_15, 0); lean_inc(x_16); -lean_dec(x_14); -x_17 = 1; -x_18 = l_Lean_Elab_Tactic_dsimpLocation_x27_go(x_1, x_15, x_17, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_16); -return x_18; +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = 1; +x_19 = l_Lean_Elab_Tactic_dsimpLocation_x27_go(x_1, x_2, x_16, x_18, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_17); +return x_19; } else { -uint8_t x_19; +uint8_t x_20; +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -3257,29 +3258,30 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_19 = !lean_is_exclusive(x_14); -if (x_19 == 0) +x_20 = !lean_is_exclusive(x_15); +if (x_20 == 0) { -return x_14; +return x_15; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_14, 0); -x_21 = lean_ctor_get(x_14, 1); +lean_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_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; +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 { -uint8_t x_23; +uint8_t x_24; +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -3289,31 +3291,32 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_23 = !lean_is_exclusive(x_11); -if (x_23 == 0) +x_24 = !lean_is_exclusive(x_12); +if (x_24 == 0) { -return x_11; +return x_12; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_11, 0); -x_25 = lean_ctor_get(x_11, 1); +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_12, 0); +x_26 = lean_ctor_get(x_12, 1); +lean_inc(x_26); lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_11); -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_dec(x_12); +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_Elab_Tactic_dsimpLocation_x27___lambda__2(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_dsimpLocation_x27___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) { _start: { -lean_object* x_13; +lean_object* x_14; +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); @@ -3321,22 +3324,22 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_4); -x_13 = l_Lean_Elab_Tactic_getFVarIds(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_13) == 0) +x_14 = l_Lean_Elab_Tactic_getFVarIds(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_14) == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_14, 0); lean_inc(x_15); -lean_dec(x_13); -x_16 = l_Lean_Elab_Tactic_dsimpLocation_x27_go(x_2, x_14, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_15); -return x_16; +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = l_Lean_Elab_Tactic_dsimpLocation_x27_go(x_2, x_3, x_15, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_16); +return x_17; } else { -uint8_t x_17; +uint8_t x_18; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -3344,65 +3347,67 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); -x_17 = !lean_is_exclusive(x_13); -if (x_17 == 0) +x_18 = !lean_is_exclusive(x_14); +if (x_18 == 0) { -return x_13; +return x_14; } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_13, 0); -x_19 = lean_ctor_get(x_13, 1); +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_14, 0); +x_20 = lean_ctor_get(x_14, 1); +lean_inc(x_20); lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_13); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_18); -lean_ctor_set(x_20, 1, x_19); -return x_20; +lean_dec(x_14); +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +return x_21; } } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_dsimpLocation_x27(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_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_dsimpLocation_x27(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, 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_2) == 0) +if (lean_obj_tag(x_3) == 0) { -lean_object* x_12; lean_object* x_13; -x_12 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_dsimpLocation_x27___lambda__1), 10, 1); -lean_closure_set(x_12, 0, x_1); -x_13 = l_Lean_Elab_Tactic_withMainContext___rarg(x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_13; +lean_object* x_13; lean_object* x_14; +x_13 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_dsimpLocation_x27___lambda__1), 11, 2); +lean_closure_set(x_13, 0, x_1); +lean_closure_set(x_13, 1, x_2); +x_14 = l_Lean_Elab_Tactic_withMainContext___rarg(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_14; } else { -lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_14 = lean_ctor_get(x_2, 0); -lean_inc(x_14); -x_15 = lean_ctor_get_uint8(x_2, sizeof(void*)*1); -lean_dec(x_2); -x_16 = lean_box(x_15); -x_17 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_dsimpLocation_x27___lambda__2___boxed), 12, 3); -lean_closure_set(x_17, 0, x_14); -lean_closure_set(x_17, 1, x_1); -lean_closure_set(x_17, 2, x_16); -x_18 = l_Lean_Elab_Tactic_withMainContext___rarg(x_17, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_18; +lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_15 = lean_ctor_get(x_3, 0); +lean_inc(x_15); +x_16 = lean_ctor_get_uint8(x_3, sizeof(void*)*1); +lean_dec(x_3); +x_17 = lean_box(x_16); +x_18 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_dsimpLocation_x27___lambda__2___boxed), 13, 4); +lean_closure_set(x_18, 0, x_15); +lean_closure_set(x_18, 1, x_1); +lean_closure_set(x_18, 2, x_2); +lean_closure_set(x_18, 3, x_17); +x_19 = l_Lean_Elab_Tactic_withMainContext___rarg(x_18, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_19; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_dsimpLocation_x27___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_dsimpLocation_x27___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) { _start: { -uint8_t x_13; lean_object* x_14; -x_13 = lean_unbox(x_3); -lean_dec(x_3); -x_14 = l_Lean_Elab_Tactic_dsimpLocation_x27___lambda__2(x_1, x_2, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_14; +uint8_t x_14; lean_object* x_15; +x_14 = lean_unbox(x_4); +lean_dec(x_4); +x_15 = l_Lean_Elab_Tactic_dsimpLocation_x27___lambda__2(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_Lean_Elab_Tactic_evalDSimpTrace___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { @@ -3438,151 +3443,155 @@ x_21 = lean_ctor_get(x_20, 0); lean_inc(x_21); if (lean_obj_tag(x_1) == 0) { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; x_22 = lean_ctor_get(x_20, 1); lean_inc(x_22); lean_dec(x_20); x_23 = lean_ctor_get(x_21, 0); lean_inc(x_23); +x_24 = lean_ctor_get(x_21, 1); +lean_inc(x_24); lean_dec(x_21); -x_24 = l_Lean_Elab_Tactic_evalSimpTrace___lambda__1___closed__2; +x_25 = l_Lean_Elab_Tactic_evalSimpTrace___lambda__1___closed__2; lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -x_25 = l_Lean_Elab_Tactic_dsimpLocation_x27(x_23, x_24, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_22); -if (lean_obj_tag(x_25) == 0) +x_26 = l_Lean_Elab_Tactic_dsimpLocation_x27(x_23, x_24, x_25, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_22); +if (lean_obj_tag(x_26) == 0) { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_26 = lean_ctor_get(x_25, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_25, 1); +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_27 = lean_ctor_get(x_26, 0); lean_inc(x_27); -lean_dec(x_25); +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +lean_dec(x_26); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -x_28 = l_Lean_Elab_Tactic_mkSimpCallStx(x_3, x_26, x_8, x_9, x_10, x_11, x_27); -x_29 = lean_ctor_get(x_28, 0); -lean_inc(x_29); -x_30 = lean_ctor_get(x_28, 1); +x_29 = l_Lean_Elab_Tactic_mkSimpCallStx(x_3, x_27, x_8, x_9, x_10, x_11, x_28); +x_30 = lean_ctor_get(x_29, 0); lean_inc(x_30); -lean_dec(x_28); -x_31 = lean_ctor_get(x_10, 5); +x_31 = lean_ctor_get(x_29, 1); lean_inc(x_31); -x_32 = l_Lean_Elab_Tactic_evalSimpTrace___lambda__2___closed__3; -x_33 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_33, 1, x_29); -x_34 = lean_box(0); -x_35 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_34); -lean_ctor_set(x_35, 3, x_34); -lean_ctor_set(x_35, 4, x_34); -lean_ctor_set(x_35, 5, x_34); -x_36 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_36, 0, x_31); -x_37 = l_Lean_Elab_Tactic_evalSimpTrace___lambda__2___closed__4; -x_38 = l_Lean_Meta_Tactic_TryThis_addSuggestion(x_2, x_35, x_36, x_37, x_34, x_8, x_9, x_10, x_11, x_30); +lean_dec(x_29); +x_32 = lean_ctor_get(x_10, 5); +lean_inc(x_32); +x_33 = l_Lean_Elab_Tactic_evalSimpTrace___lambda__2___closed__3; +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_30); +x_35 = lean_box(0); +x_36 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +lean_ctor_set(x_36, 2, x_35); +lean_ctor_set(x_36, 3, x_35); +lean_ctor_set(x_36, 4, x_35); +lean_ctor_set(x_36, 5, x_35); +x_37 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_37, 0, x_32); +x_38 = l_Lean_Elab_Tactic_evalSimpTrace___lambda__2___closed__4; +x_39 = l_Lean_Meta_Tactic_TryThis_addSuggestion(x_2, x_36, x_37, x_38, x_35, x_8, x_9, x_10, x_11, x_31); lean_dec(x_9); lean_dec(x_8); -return x_38; +return x_39; } else { -uint8_t x_39; +uint8_t x_40; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); -x_39 = !lean_is_exclusive(x_25); -if (x_39 == 0) +x_40 = !lean_is_exclusive(x_26); +if (x_40 == 0) { -return x_25; +return x_26; } else { -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_25, 0); -x_41 = lean_ctor_get(x_25, 1); +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_26, 0); +x_42 = lean_ctor_get(x_26, 1); +lean_inc(x_42); lean_inc(x_41); -lean_inc(x_40); -lean_dec(x_25); -x_42 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_42, 0, x_40); -lean_ctor_set(x_42, 1, x_41); -return x_42; +lean_dec(x_26); +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_43; lean_object* x_44; uint8_t x_45; -x_43 = lean_ctor_get(x_20, 1); -lean_inc(x_43); -lean_dec(x_20); -x_44 = lean_ctor_get(x_21, 0); +lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; +x_44 = lean_ctor_get(x_20, 1); lean_inc(x_44); +lean_dec(x_20); +x_45 = lean_ctor_get(x_21, 0); +lean_inc(x_45); +x_46 = lean_ctor_get(x_21, 1); +lean_inc(x_46); lean_dec(x_21); -x_45 = !lean_is_exclusive(x_1); -if (x_45 == 0) +x_47 = !lean_is_exclusive(x_1); +if (x_47 == 0) { -lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_46 = lean_ctor_get(x_1, 0); -x_47 = l_Lean_Elab_Tactic_expandLocation(x_46); -lean_dec(x_46); +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_1, 0); +x_49 = l_Lean_Elab_Tactic_expandLocation(x_48); +lean_dec(x_48); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -x_48 = l_Lean_Elab_Tactic_dsimpLocation_x27(x_44, x_47, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_43); -if (lean_obj_tag(x_48) == 0) +x_50 = l_Lean_Elab_Tactic_dsimpLocation_x27(x_45, x_46, x_49, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_44); +if (lean_obj_tag(x_50) == 0) { -lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_49 = lean_ctor_get(x_48, 0); -lean_inc(x_49); -x_50 = lean_ctor_get(x_48, 1); -lean_inc(x_50); -lean_dec(x_48); +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_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); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -x_51 = l_Lean_Elab_Tactic_mkSimpCallStx(x_3, x_49, x_8, x_9, x_10, x_11, x_50); -x_52 = lean_ctor_get(x_51, 0); -lean_inc(x_52); -x_53 = lean_ctor_get(x_51, 1); -lean_inc(x_53); -lean_dec(x_51); -x_54 = lean_ctor_get(x_10, 5); +x_53 = l_Lean_Elab_Tactic_mkSimpCallStx(x_3, x_51, x_8, x_9, x_10, x_11, x_52); +x_54 = lean_ctor_get(x_53, 0); lean_inc(x_54); -x_55 = l_Lean_Elab_Tactic_evalSimpTrace___lambda__2___closed__3; -x_56 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_56, 0, x_55); -lean_ctor_set(x_56, 1, x_52); -x_57 = lean_box(0); -x_58 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_58, 0, x_56); -lean_ctor_set(x_58, 1, x_57); -lean_ctor_set(x_58, 2, x_57); -lean_ctor_set(x_58, 3, x_57); -lean_ctor_set(x_58, 4, x_57); -lean_ctor_set(x_58, 5, x_57); -lean_ctor_set(x_1, 0, x_54); -x_59 = l_Lean_Elab_Tactic_evalSimpTrace___lambda__2___closed__4; -x_60 = l_Lean_Meta_Tactic_TryThis_addSuggestion(x_2, x_58, x_1, x_59, x_57, x_8, x_9, x_10, x_11, x_53); +x_55 = lean_ctor_get(x_53, 1); +lean_inc(x_55); +lean_dec(x_53); +x_56 = lean_ctor_get(x_10, 5); +lean_inc(x_56); +x_57 = l_Lean_Elab_Tactic_evalSimpTrace___lambda__2___closed__3; +x_58 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_58, 0, x_57); +lean_ctor_set(x_58, 1, x_54); +x_59 = lean_box(0); +x_60 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_60, 0, x_58); +lean_ctor_set(x_60, 1, x_59); +lean_ctor_set(x_60, 2, x_59); +lean_ctor_set(x_60, 3, x_59); +lean_ctor_set(x_60, 4, x_59); +lean_ctor_set(x_60, 5, x_59); +lean_ctor_set(x_1, 0, x_56); +x_61 = l_Lean_Elab_Tactic_evalSimpTrace___lambda__2___closed__4; +x_62 = l_Lean_Meta_Tactic_TryThis_addSuggestion(x_2, x_60, x_1, x_61, x_59, x_8, x_9, x_10, x_11, x_55); lean_dec(x_9); lean_dec(x_8); -return x_60; +return x_62; } else { -uint8_t x_61; +uint8_t x_63; lean_free_object(x_1); lean_dec(x_11); lean_dec(x_10); @@ -3590,115 +3599,115 @@ lean_dec(x_9); lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); -x_61 = !lean_is_exclusive(x_48); -if (x_61 == 0) +x_63 = !lean_is_exclusive(x_50); +if (x_63 == 0) { -return x_48; +return x_50; } else { -lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_62 = lean_ctor_get(x_48, 0); -x_63 = lean_ctor_get(x_48, 1); -lean_inc(x_63); -lean_inc(x_62); -lean_dec(x_48); -x_64 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_64, 0, x_62); -lean_ctor_set(x_64, 1, x_63); -return x_64; +lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_64 = lean_ctor_get(x_50, 0); +x_65 = lean_ctor_get(x_50, 1); +lean_inc(x_65); +lean_inc(x_64); +lean_dec(x_50); +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_65; lean_object* x_66; lean_object* x_67; -x_65 = lean_ctor_get(x_1, 0); -lean_inc(x_65); +lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_1, 0); +lean_inc(x_67); lean_dec(x_1); -x_66 = l_Lean_Elab_Tactic_expandLocation(x_65); -lean_dec(x_65); +x_68 = l_Lean_Elab_Tactic_expandLocation(x_67); +lean_dec(x_67); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -x_67 = l_Lean_Elab_Tactic_dsimpLocation_x27(x_44, x_66, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_43); -if (lean_obj_tag(x_67) == 0) +x_69 = l_Lean_Elab_Tactic_dsimpLocation_x27(x_45, x_46, x_68, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_44); +if (lean_obj_tag(x_69) == 0) { -lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; -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_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_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_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -x_70 = l_Lean_Elab_Tactic_mkSimpCallStx(x_3, x_68, x_8, x_9, x_10, x_11, x_69); -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_10, 5); +x_72 = l_Lean_Elab_Tactic_mkSimpCallStx(x_3, x_70, x_8, x_9, x_10, x_11, x_71); +x_73 = lean_ctor_get(x_72, 0); lean_inc(x_73); -x_74 = l_Lean_Elab_Tactic_evalSimpTrace___lambda__2___closed__3; -x_75 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_75, 0, x_74); -lean_ctor_set(x_75, 1, x_71); -x_76 = lean_box(0); -x_77 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_77, 0, x_75); -lean_ctor_set(x_77, 1, x_76); -lean_ctor_set(x_77, 2, x_76); -lean_ctor_set(x_77, 3, x_76); -lean_ctor_set(x_77, 4, x_76); -lean_ctor_set(x_77, 5, x_76); -x_78 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_78, 0, x_73); -x_79 = l_Lean_Elab_Tactic_evalSimpTrace___lambda__2___closed__4; -x_80 = l_Lean_Meta_Tactic_TryThis_addSuggestion(x_2, x_77, x_78, x_79, x_76, x_8, x_9, x_10, x_11, x_72); +x_74 = lean_ctor_get(x_72, 1); +lean_inc(x_74); +lean_dec(x_72); +x_75 = lean_ctor_get(x_10, 5); +lean_inc(x_75); +x_76 = l_Lean_Elab_Tactic_evalSimpTrace___lambda__2___closed__3; +x_77 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_77, 0, x_76); +lean_ctor_set(x_77, 1, x_73); +x_78 = lean_box(0); +x_79 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_79, 0, x_77); +lean_ctor_set(x_79, 1, x_78); +lean_ctor_set(x_79, 2, x_78); +lean_ctor_set(x_79, 3, x_78); +lean_ctor_set(x_79, 4, x_78); +lean_ctor_set(x_79, 5, x_78); +x_80 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_80, 0, x_75); +x_81 = l_Lean_Elab_Tactic_evalSimpTrace___lambda__2___closed__4; +x_82 = l_Lean_Meta_Tactic_TryThis_addSuggestion(x_2, x_79, x_80, x_81, x_78, x_8, x_9, x_10, x_11, x_74); lean_dec(x_9); lean_dec(x_8); -return x_80; +return x_82; } else { -lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); -x_81 = lean_ctor_get(x_67, 0); -lean_inc(x_81); -x_82 = lean_ctor_get(x_67, 1); -lean_inc(x_82); -if (lean_is_exclusive(x_67)) { - lean_ctor_release(x_67, 0); - lean_ctor_release(x_67, 1); - x_83 = x_67; +x_83 = lean_ctor_get(x_69, 0); +lean_inc(x_83); +x_84 = lean_ctor_get(x_69, 1); +lean_inc(x_84); +if (lean_is_exclusive(x_69)) { + lean_ctor_release(x_69, 0); + lean_ctor_release(x_69, 1); + x_85 = x_69; } else { - lean_dec_ref(x_67); - x_83 = lean_box(0); + lean_dec_ref(x_69); + x_85 = lean_box(0); } -if (lean_is_scalar(x_83)) { - x_84 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_85)) { + x_86 = lean_alloc_ctor(1, 2, 0); } else { - x_84 = x_83; + x_86 = x_85; } -lean_ctor_set(x_84, 0, x_81); -lean_ctor_set(x_84, 1, x_82); -return x_84; +lean_ctor_set(x_86, 0, x_83); +lean_ctor_set(x_86, 1, x_84); +return x_86; } } } } else { -uint8_t x_85; +uint8_t x_87; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -3710,23 +3719,23 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_85 = !lean_is_exclusive(x_20); -if (x_85 == 0) +x_87 = !lean_is_exclusive(x_20); +if (x_87 == 0) { return x_20; } else { -lean_object* x_86; lean_object* x_87; lean_object* x_88; -x_86 = lean_ctor_get(x_20, 0); -x_87 = lean_ctor_get(x_20, 1); -lean_inc(x_87); -lean_inc(x_86); +lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_88 = lean_ctor_get(x_20, 0); +x_89 = lean_ctor_get(x_20, 1); +lean_inc(x_89); +lean_inc(x_88); lean_dec(x_20); -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; +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; } } } @@ -4546,7 +4555,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalDSimpTrace_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(79u); +x_1 = lean_unsigned_to_nat(80u); x_2 = lean_unsigned_to_nat(29u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -4558,7 +4567,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalDSimpTrace_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(90u); +x_1 = lean_unsigned_to_nat(92u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -4586,7 +4595,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalDSimpTrace_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(79u); +x_1 = lean_unsigned_to_nat(80u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -4598,7 +4607,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalDSimpTrace_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(79u); +x_1 = lean_unsigned_to_nat(80u); x_2 = lean_unsigned_to_nat(47u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); diff --git a/stage0/stdlib/Lean/Environment.c b/stage0/stdlib/Lean/Environment.c index d2cb24b63220..4694d3819ce9 100644 --- a/stage0/stdlib/Lean/Environment.c +++ b/stage0/stdlib/Lean/Environment.c @@ -17,8 +17,8 @@ LEAN_EXPORT uint32_t l_Lean_EnvironmentHeader_trustLevel___default; uint8_t lean_compacted_region_is_memory_mapped(size_t); LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_mkMapDeclarationExtension___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_ConstantInfo_instantiateTypeLevelParams(lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_2443____closed__3; LEAN_EXPORT lean_object* l_Lean_throwAlreadyImported___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6189____spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_displayStats___spec__8___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Environment_find_x3f___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Init_Util_0__outOfBounds___rarg(lean_object*); @@ -50,12 +50,9 @@ static lean_object* l_Lean_Environment_displayStats___closed__10; LEAN_EXPORT lean_object* l_Lean_registerPersistentEnvExtensionUnsafe(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries(lean_object*, lean_object*); lean_object* lean_uint32_to_nat(uint32_t); -LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6189____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_mkEmptyEnvironment___lambda__1___closed__3; -static lean_object* l___auto____x40_Lean_Environment___hyg_2418____closed__8; LEAN_EXPORT lean_object* l_Lean_importModules___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4___closed__1; -static lean_object* l___auto____x40_Lean_Environment___hyg_2418____closed__9; static lean_object* l_Lean_registerPersistentEnvExtensionUnsafe___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_instInhabitedModuleData; LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_modifyState___rarg___boxed(lean_object*, lean_object*, lean_object*); @@ -75,9 +72,9 @@ static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_ LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___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*); static lean_object* l_List_toString___at_Lean_Environment_displayStats___spec__1___closed__1; LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_mkTagDeclarationExtension___spec__1___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6189____closed__4; uint8_t lean_usize_dec_le(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_EnvironmentHeader_regions___default; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MapDeclarationExtension_insert(lean_object*); lean_object* l_Lean_ConstantInfo_levelParams(lean_object*); static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__8; @@ -88,10 +85,11 @@ LEAN_EXPORT lean_object* l_Lean_HashMap_numBuckets___at_Lean_Environment_display LEAN_EXPORT lean_object* l_Lean_mkMapDeclarationExtension___rarg___lambda__2___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_setImportedEntries(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6189____closed__2; +static lean_object* l___auto____x40_Lean_Environment___hyg_2443____closed__14; size_t lean_uint64_to_usize(uint64_t); LEAN_EXPORT lean_object* l_Lean_finalizeImport___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_contains___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__7___boxed(lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_2443____closed__28; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_finalizeImport___spec__8___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_EnvironmentHeader_quotInit___default; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__4___lambda__1___boxed(lean_object*, lean_object*, lean_object*); @@ -113,7 +111,8 @@ LEAN_EXPORT lean_object* l_Lean_AssocList_contains___at_Lean_finalizeImport___sp lean_object* lean_kernel_is_def_eq(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___auto____x40_Lean_Environment___hyg_2418____closed__10; +static lean_object* l___auto____x40_Lean_Environment___hyg_2443____closed__13; +static lean_object* l___auto____x40_Lean_Environment___hyg_2443____closed__4; lean_object* l_Array_qpartition___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkHashSetImp___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_getModuleEntries(lean_object*, lean_object*, lean_object*); @@ -127,9 +126,8 @@ static lean_object* l_Lean_instInhabitedEnvExtensionInterface___closed__4; lean_object* lean_mk_array(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_displayStats___spec__9(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkExtNameMap(lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2418____closed__22; +static lean_object* l___auto____x40_Lean_Environment___hyg_2443____closed__26; LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtensionDescr_statsFn___default(lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2418____closed__14; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_mkModuleData___spec__1(lean_object*, size_t, size_t, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_importModulesCore(lean_object*, lean_object*, lean_object*); @@ -142,7 +140,7 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Environmen LEAN_EXPORT lean_object* l_Lean_readModuleData___boxed(lean_object*, lean_object*); static lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4___closed__2; LEAN_EXPORT uint8_t l_Lean_Environment_isConstructor(lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6189____lambda__2___closed__1; +static lean_object* l___auto____x40_Lean_Environment___hyg_2443____closed__18; LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_modifyState(lean_object*, lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_getState(lean_object*, lean_object*, lean_object*); @@ -155,33 +153,34 @@ static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_ LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_EnvExtensionInterfaceUnsafe_envExtensionsRef; LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvExtension_modifyState___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_2443____closed__9; lean_object* lean_get_num_attributes(lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvironmentHeader_moduleNames___default; LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_mkInitialExtStates(lean_object*); LEAN_EXPORT lean_object* lean_environment_find(lean_object*, lean_object*); static lean_object* l_Lean_TagDeclarationExtension_instInhabitedTagDeclarationExtension___closed__1; lean_object* l_Lean_RBNode_insert___at_Lean_NameSet_insert___spec__1(lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_2443____closed__22; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3(lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMapImp_expand___at_Lean_mkExtNameMap___spec__4(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__4___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_Environment___hyg_6214____lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Environment_imports___boxed(lean_object*); static lean_object* l_List_toString___at_Lean_Environment_displayStats___spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtensionDescr_toArrayFn___default___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMapImp_find_x3f___at_Lean_Environment_find_x3f___spec__5(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkModuleData___lambda__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getNumBuiltinAttributes___boxed(lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2418____closed__1; static lean_object* l_List_toString___at_Lean_Environment_displayStats___spec__1___closed__3; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_mkModuleData___spec__3___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2418____closed__23; -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6189_(lean_object*); +LEAN_EXPORT lean_object* l___auto____x40_Lean_Environment___hyg_2846_; LEAN_EXPORT lean_object* l_Lean_Environment_addDecl___boxed(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__2; -LEAN_EXPORT lean_object* l___auto____x40_Lean_Environment___hyg_3160_; LEAN_EXPORT lean_object* l_Lean_ImportStateM_run(lean_object*); LEAN_EXPORT lean_object* l_Lean_importModules___lambda__1(lean_object*, lean_object*, uint32_t, uint8_t, lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__6; uint8_t l_Lean_Name_quickLt(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6189____lambda__2(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_ConstantInfo_isUnsafe(lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMap_insert___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__6(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__2___spec__1(lean_object*, lean_object*); @@ -193,7 +192,6 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Environm LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__5(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_addExtraName(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_isNamespace___boxed(lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2418____closed__24; LEAN_EXPORT lean_object* l_Lean_mkTagDeclarationExtension___lambda__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Environment_find_x3f___spec__3___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_MapDeclarationExtension_contains___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -202,8 +200,9 @@ LEAN_EXPORT lean_object* l_Lean_updateEnvAttributes___boxed(lean_object*, lean_o LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_importModules___closed__1; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_mkModuleData___spec__4___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2418____closed__20; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_System_FilePath_pathExists(lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_2443____closed__15; LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___lambda__1(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_75_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMapImp_contains___at_Lean_Environment_addExtraName___spec__2___boxed(lean_object*, lean_object*); @@ -216,16 +215,14 @@ LEAN_EXPORT lean_object* l_Lean_EnvironmentHeader_imports___default; lean_object* lean_string_push(lean_object*, uint32_t); LEAN_EXPORT lean_object* l_Lean_Environment_registerNamespace(lean_object*, lean_object*); uint8_t l_Std_Format_isNil(lean_object*); -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6189____lambda__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6189____spec__3(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_modifyState___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_mkModuleData___spec__5___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_mkModuleData___spec__3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_setState___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_SMap_size___at_Lean_Environment_displayStats___spec__3(lean_object*); +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_initFn____x40_Lean_Environment___hyg_1314_(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_finalizeImport___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_throwAlreadyImported___spec__1(lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2418____closed__17; lean_object* l_Lean_initializing(lean_object*); extern lean_object* l_Lean_NameSet_instInhabitedNameSet; LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension(lean_object*, lean_object*, lean_object*); @@ -247,7 +244,6 @@ LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_mkEmptyEnvironment___spec__1 LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_getState___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_TagDeclarationExtension_isTagged(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_contains___at_Lean_mkExtNameMap___spec__3___boxed(lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2418____closed__28; LEAN_EXPORT lean_object* l_Lean_ImportState_moduleNames___default; uint8_t l_Lean_HashSetImp_contains___at_Lean_NameHashSet_contains___spec__1(lean_object*, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_displayStats___spec__8___lambda__1___closed__2; @@ -266,14 +262,15 @@ LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_setState___rarg___l LEAN_EXPORT lean_object* l_Lean_Environment_evalConstCheck(lean_object*); static lean_object* l_Lean_throwAlreadyImported___rarg___closed__5; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_2443____closed__8; LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_setState(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_mkModuleData___spec__1___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_EnvExtensionInterfaceUnsafe_getState___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2418____closed__3; LEAN_EXPORT lean_object* l_Lean_instToStringImport(lean_object*); LEAN_EXPORT uint8_t l_Array_qsort_sort___at_Lean_mkMapDeclarationExtension___spec__1___rarg___lambda__1(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_registerPersistentEnvExtensionUnsafe___spec__1___rarg(lean_object*, lean_object*, size_t, size_t); +static lean_object* l___auto____x40_Lean_Environment___hyg_2443____closed__2; LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_ensureExtensionsArraySize_loop(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_finalizeImport___lambda__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_addEntry___rarg___lambda__1(lean_object*, lean_object*, lean_object*); @@ -282,6 +279,7 @@ LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_getModuleEntries___rarg(l static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__2; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_freeRegions___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*); lean_object* lean_nat_to_int(lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_2443____closed__6; LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_Environment_throwUnexpectedType___rarg(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__1___closed__1; lean_object* lean_nat_div(lean_object*, lean_object*); @@ -290,7 +288,6 @@ static lean_object* l_Lean_mkMapDeclarationExtension___rarg___closed__2; static lean_object* l_Lean_Environment_displayStats___closed__7; static lean_object* l_Lean_Environment_header___default___closed__1; static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg___closed__1; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6189____spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_AssocList_contains___at_Lean_finalizeImport___spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_75____closed__24; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_finalizeImport___spec__8(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); @@ -321,8 +318,8 @@ LEAN_EXPORT lean_object* l_Lean_Kernel_isDefEq___boxed(lean_object*, lean_object LEAN_EXPORT lean_object* lean_display_stats(lean_object*, lean_object*); uint8_t l_Lean_NameMap_contains___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvExtension_modifyState___rarg___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2418____closed__18; static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___closed__1; +static lean_object* l___auto____x40_Lean_Environment___hyg_2443____closed__10; static lean_object* l_Lean_Environment_displayStats___closed__8; LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__1(lean_object*); static size_t l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3___closed__2; @@ -334,8 +331,8 @@ LEAN_EXPORT uint8_t l_Array_binSearchAux___at_Lean_MapDeclarationExtension_conta lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instMonadEnv___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension___lambda__3___boxed(lean_object*); +LEAN_EXPORT lean_object* l___auto____x40_Lean_Environment___hyg_3185_; LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_Environment_registerNamePrefixes(lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6189____closed__6; LEAN_EXPORT uint8_t l_Array_binSearchAux___at_Lean_TagDeclarationExtension_isTagged___spec__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Environment_displayStats___closed__4; LEAN_EXPORT lean_object* l_Lean_instInhabitedEnvExtensionInterface___lambda__2(lean_object*, lean_object*, lean_object*); @@ -352,6 +349,7 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_mkExtNameMap___spe lean_object* lean_list_to_array(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_binSearchAux___at_Lean_MapDeclarationExtension_find_x3f___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withImportModules___rarg(lean_object*, lean_object*, uint32_t, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__4(lean_object*); lean_object* l_EStateM_pure___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtensionState(lean_object*, lean_object*); static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_75____closed__10; @@ -360,19 +358,19 @@ static lean_object* l_Lean_throwAlreadyImported___rarg___closed__7; static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_75____closed__4; LEAN_EXPORT lean_object* l_Lean_instInhabitedEnvExtensionInterface___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Environment_displayStats___closed__9; +static lean_object* l___auto____x40_Lean_Environment___hyg_2443____closed__11; +LEAN_EXPORT lean_object* l___auto____x40_Lean_Environment___hyg_2443_; 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_RBTree_toArray___at_Lean_mkModuleData___spec__7(lean_object*); LEAN_EXPORT lean_object* l_Lean_ImportState_moduleNameSet___default; LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_setState(lean_object*); LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_setState___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2418____closed__13; LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at_Lean_Environment_addExtraName___spec__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* lean_environment_set_main_module(lean_object*, lean_object*); static lean_object* l_Lean_TagDeclarationExtension_tag___closed__3; static lean_object* l_Lean_instToStringImport___closed__1; LEAN_EXPORT lean_object* l_Lean_HashMap_insert___at_Lean_mkExtNameMap___spec__2(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_getState___rarg___closed__1; -LEAN_EXPORT lean_object* l___auto____x40_Lean_Environment___hyg_2418_; lean_object* l_List_lengthTRAux___rarg(lean_object*, lean_object*); LEAN_EXPORT uint32_t lean_environment_trust_level(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_getModuleEntries___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -381,12 +379,11 @@ LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg___la uint8_t lean_name_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_SMap_size___at_Lean_Environment_displayStats___spec__3___boxed(lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2418____closed__15; LEAN_EXPORT lean_object* l_Lean_MapDeclarationExtension_find_x3f___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_2443____closed__20; LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_MapDeclarationExtension_find_x3f___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_modifyState___rarg___lambda__1(lean_object*, lean_object*); lean_object* l_instToStringNat(lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2418____closed__25; LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_modifyState(lean_object*, lean_object*); static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_75____closed__23; LEAN_EXPORT lean_object* l_Lean_ConstantInfo_instantiateValueLevelParams_x21(lean_object*, lean_object*); @@ -409,7 +406,6 @@ static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_ LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_Environment_isNamespaceName___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_ConstantInfo_instantiateValueLevelParams_x21___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_75____closed__14; -static lean_object* l___auto____x40_Lean_Environment___hyg_2418____closed__12; LEAN_EXPORT lean_object* l_panic___at_Lean_MapDeclarationExtension_insert___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_instToStringImport___closed__2; static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_75____closed__5; @@ -419,7 +415,6 @@ LEAN_EXPORT lean_object* l_Lean_Environment_getModuleIdx_x3f___lambda__1___boxed LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_mkMapDeclarationExtension___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_instantiateLevelParams(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension___lambda__4(lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2418____closed__5; LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__2(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_TagDeclarationExtension_tag___closed__2; LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___rarg(lean_object*, lean_object*, lean_object*); @@ -429,8 +424,11 @@ LEAN_EXPORT lean_object* l_Lean_instInhabitedImport; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_EnvExtensionInterfaceUnsafe_mkInitialExtStates___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_ensureExtensionsArraySize(lean_object*, lean_object*); static lean_object* l_Lean_mkEmptyEnvironment___closed__2; +static lean_object* l___auto____x40_Lean_Environment___hyg_2443____closed__27; LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__3(lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_2443____closed__5; extern lean_object* l_Std_Format_defWidth; +LEAN_EXPORT lean_object* l___auto____x40_Lean_Environment___hyg_3363_; LEAN_EXPORT lean_object* l_Lean_AssocList_replace___at_Lean_mkExtNameMap___spec__7(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Environment_hasUnsafe(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_getEntries___rarg(lean_object*, lean_object*, lean_object*); @@ -441,6 +439,7 @@ static lean_object* l_Lean_Environment_displayStats___closed__6; LEAN_EXPORT lean_object* l_Lean_registerEnvExtension___rarg(lean_object*, lean_object*); static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_75____closed__19; LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_Environment_getModuleIdxFor_x3f___spec__2(lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__1; lean_object* lean_usize_to_nat(size_t); size_t lean_hashmap_mk_idx(lean_object*, uint64_t); static lean_object* l_Lean_Environment_evalConstCheck___rarg___closed__1; @@ -457,6 +456,7 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_importModules___spec LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Environment_find_x3f___spec__3(lean_object*, size_t, lean_object*); static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_75____closed__2; LEAN_EXPORT lean_object* l_Lean_throwAlreadyImported___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__2(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_mkEmptyEnvironment___spec__2(lean_object*); static lean_object* l_Lean_instInhabitedEnvExtensionInterface___closed__1; LEAN_EXPORT lean_object* l_Lean_HashMapImp_moveEntries___at_Lean_finalizeImport___spec__4(lean_object*, lean_object*, lean_object*); @@ -470,9 +470,9 @@ LEAN_EXPORT lean_object* l_List_toString___at_Lean_Environment_displayStats___sp static lean_object* l_Lean_Environment_registerNamespace___closed__1; LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_mkTagDeclarationExtension___closed__2; -static lean_object* l___auto____x40_Lean_Environment___hyg_2418____closed__19; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_mkModuleData___spec__5___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_EnvExtensionStateSpec___closed__1; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6214____lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* lean_environment_add(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Environment_getModuleIdx_x3f___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkModuleData___lambda__2(lean_object*, lean_object*, lean_object*); @@ -480,7 +480,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_EnvExtensionInterf lean_object* l_Lean_Expr_FindImpl_findUnsafe_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_displayStats___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_2376_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_hasUnsafe___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_getEntries(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_finalizeImport___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -495,6 +494,7 @@ LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_setState___rarg___lambda_ static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_75____closed__18; LEAN_EXPORT lean_object* l_Lean_Environment_getModuleIdx_x3f___boxed(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_importModules___spec__1___closed__1; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__3(lean_object*, size_t, size_t, lean_object*); extern lean_object* l_Lean_NameSet_empty; static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_modifyState___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_mkMapDeclarationExtension___rarg___lambda__2(lean_object*); @@ -503,6 +503,7 @@ LEAN_EXPORT lean_object* l_Lean_SMap_stageSizes___at_Lean_Environment_displaySta static lean_object* l_Lean_MapDeclarationExtension_instInhabitedMapDeclarationExtension___closed__1; lean_object* lean_string_length(lean_object*); static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__6; +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__3; static lean_object* l_Lean_instReprImport___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___rarg___lambda__1(lean_object*, lean_object*, lean_object*); @@ -513,19 +514,19 @@ LEAN_EXPORT lean_object* l_Lean_instReprImport; LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_displayStats___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_finalizeImport___spec__9(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_2443____closed__19; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_getState(lean_object*, lean_object*); LEAN_EXPORT lean_object* lean_environment_main_module(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_displayStats___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2418____closed__21; LEAN_EXPORT lean_object* l_Lean_EnvExtension_setState___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_replace___at_Lean_finalizeImport___spec__6(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Environment_displayStats___closed__1; static lean_object* l_Lean_instInhabitedPersistentEnvExtension___closed__3; LEAN_EXPORT lean_object* l_Lean_EnvExtension_getState(lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2418____closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__4(size_t, 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_initFn____x40_Lean_Environment___hyg_6214____closed__4; LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_ensureExtensionsArraySize_loop___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtensionDescr_statsFn___default___boxed(lean_object*, lean_object*); @@ -548,6 +549,7 @@ LEAN_EXPORT lean_object* l_Lean_registerPersistentEnvExtensionUnsafe___rarg___la LEAN_EXPORT uint8_t l_Lean_SMap_contains___at_Lean_Environment_addExtraName___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ConstantInfo_instantiateTypeLevelParams___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_TagDeclarationExtension_tag(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6214_(lean_object*); static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__5; uint8_t l_Lean_NameSet_contains(lean_object*, lean_object*); static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_75____closed__8; @@ -556,7 +558,6 @@ uint64_t l_Lean_Name_hash___override(lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedEnvExtensionInterface___lambda__1(lean_object*, lean_object*); static lean_object* l_Lean_mkTagDeclarationExtension___closed__1; LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_MapDeclarationExtension_find_x3f___spec__1___rarg___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6189____spec__4(lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_EnvExtensionInterfaceUnsafe_setState___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension___lambda__2___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -566,16 +567,19 @@ LEAN_EXPORT lean_object* l_Lean_MapDeclarationExtension_contains___rarg___boxed( static lean_object* l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___closed__1; static lean_object* l_Lean_throwAlreadyImported___rarg___closed__3; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__4___lambda__1(lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_2443____closed__21; lean_object* lean_nat_sub(lean_object*, lean_object*); static lean_object* l_Lean_throwAlreadyImported___rarg___closed__6; LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension___lambda__2(lean_object*, lean_object*); uint32_t lean_uint32_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_mkModuleData___spec__2(lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_2443____closed__12; LEAN_EXPORT lean_object* l_Lean_instInhabitedEnvExtensionInterface; static lean_object* l_Lean_instInhabitedPersistentEnvExtension___closed__5; lean_object* lean_nat_mul(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Kernel_isDefEqGuarded(lean_object*, lean_object*, lean_object*, lean_object*); static uint32_t l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___lambda__1___closed__1; +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__2; LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__1___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkMapDeclarationExtension___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_getState___rarg(lean_object*, lean_object*, lean_object*); @@ -595,12 +599,10 @@ LEAN_EXPORT lean_object* l_panic___at_Lean_EnvExtensionInterfaceUnsafe_modifySta lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_CompactedRegion_free___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_Environment_getModuleIdxFor_x3f___spec__2___boxed(lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2418____closed__4; LEAN_EXPORT lean_object* l_Lean_Kernel_whnf___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_getState___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_getModuleIdxFor_x3f(lean_object*, lean_object*); static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__7; -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6189____closed__5; LEAN_EXPORT lean_object* l_Lean_persistentEnvExtensionsRef; lean_object* lean_io_initializing(lean_object*); lean_object* l_List_reverse___rarg(lean_object*); @@ -613,19 +615,19 @@ size_t lean_usize_sub(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtensionState___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedEnvExtensionInterface___lambda__1___boxed(lean_object*, lean_object*); lean_object* l_Lean_mkHashMapImp___rarg(lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2418____closed__7; +static lean_object* l___auto____x40_Lean_Environment___hyg_2443____closed__16; LEAN_EXPORT lean_object* l_Lean_namespacesExt; LEAN_EXPORT lean_object* l_Lean_EnvExtension_modifyState(lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_mkExtNameMap___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedPersistentEnvExtension___closed__1; LEAN_EXPORT lean_object* l_Lean_TagDeclarationExtension_isTagged___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_2443____closed__17; uint8_t l_Lean_Name_quickCmp(lean_object*, lean_object*); static size_t l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_75____boxed(lean_object*, lean_object*); static lean_object* l_Lean_mkMapDeclarationExtension___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_importModulesCore___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ModuleIdx_toNat(lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2418____closed__16; LEAN_EXPORT lean_object* l_Lean_HashMapImp_moveEntries___at_Lean_mkExtNameMap___spec__5(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__2(lean_object*); size_t lean_usize_add(size_t, size_t); @@ -640,8 +642,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_Environment_isQuot LEAN_EXPORT lean_object* l_Lean_instInhabitedEnvExtensionInterface___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); static lean_object* l_Lean_MapDeclarationExtension_find_x3f___rarg___closed__1; -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6189____closed__3; -LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_initFn____x40_Lean_Environment___hyg_1289_(lean_object*); static lean_object* l_Lean_Environment_displayStats___closed__5; lean_object* l_Array_foldlMUnsafe_fold___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__1; @@ -658,23 +658,25 @@ LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_modifyState___rarg( LEAN_EXPORT lean_object* l_Lean_RBNode_fold___at_Lean_mkModuleData___spec__8(lean_object*, 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*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2418____closed__27; lean_object* l_List_redLength___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_Environment_find_x3f___spec__6(lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__5; LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__3___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_2443____closed__23; static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_75____closed__22; LEAN_EXPORT uint8_t lean_environment_quot_init(lean_object*); static lean_object* l_Lean_throwAlreadyImported___rarg___closed__4; LEAN_EXPORT lean_object* l_Lean_HashMapImp_expand___at_Lean_finalizeImport___spec__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_mkExtNameMap___spec__1(lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6189____closed__1; LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_finalizeImport___spec__5(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_SMap_stageSizes___at_Lean_Environment_displayStats___spec__4(lean_object*); static lean_object* l_Lean_mkModuleData___closed__2; LEAN_EXPORT lean_object* l_Lean_Environment_getNamespaceSet___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMapImp_find_x3f___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__2(lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_2443____closed__24; static lean_object* l_Lean_mkEmptyEnvironment___lambda__1___closed__1; lean_object* lean_string_append(lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_2443____closed__1; LEAN_EXPORT lean_object* l_Lean_instMonadEnv(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_evalConstCheck___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__2(lean_object*, lean_object*, lean_object*); @@ -686,20 +688,19 @@ static lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg___closed_ lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_Array_binSearchAux___at_Lean_MapDeclarationExtension_contains___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2418____closed__26; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_2401_(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_mkExtNameMap___spec__1___boxed(lean_object*); static lean_object* l_Lean_Environment_displayStats___closed__2; LEAN_EXPORT lean_object* l_Lean_saveModuleData___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2418____closed__11; uint8_t lean_nat_dec_le(lean_object*, lean_object*); lean_object* l_Lean_SMap_insert___at_Lean_NameSSet_insert___spec__1(lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); LEAN_EXPORT lean_object* lean_mk_empty_environment(uint32_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkTagDeclarationExtension___lambda__1(lean_object*, lean_object*); static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___closed__2; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6189____spec__2(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkMapDeclarationExtension___boxed(lean_object*, lean_object*); static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_getState___rarg___closed__2; +static lean_object* l___auto____x40_Lean_Environment___hyg_2443____closed__25; lean_object* lean_compacted_region_free(size_t, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__5___boxed(lean_object*, lean_object*); @@ -714,9 +715,8 @@ LEAN_EXPORT lean_object* lean_import_modules(lean_object*, lean_object*, uint32_ LEAN_EXPORT lean_object* l_Lean_HashMapImp_moveEntries___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__9(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_registerExt(lean_object*); LEAN_EXPORT lean_object* l_Lean_importModules___boxed__const__1; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6189____spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__1(lean_object*, lean_object*); lean_object* lean_kernel_whnf(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___auto____x40_Lean_Environment___hyg_3338_; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_mkEmptyEnvironment___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_getState___rarg___boxed(lean_object*, lean_object*, lean_object*); @@ -728,7 +728,6 @@ LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg___la LEAN_EXPORT lean_object* l_Lean_mkMapDeclarationExtension(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_setState___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MapDeclarationExtension_contains(lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6189____spec__5(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvironmentHeader_mainModule___default; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* lean_environment_free_regions(lean_object*, lean_object*); @@ -740,17 +739,17 @@ static lean_object* l_List_foldl___at_Lean_Environment_displayStats___spec__2___ static lean_object* l_Lean_instInhabitedPersistentEnvExtension___closed__4; LEAN_EXPORT lean_object* l_Lean_SMap_numBuckets___at_Lean_Environment_displayStats___spec__5___boxed(lean_object*); lean_object* l_Nat_repr(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__5(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Environment_addExtraName___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvExtension_setState(lean_object*); -LEAN_EXPORT lean_object* l___auto____x40_Lean_Environment___hyg_2821_; LEAN_EXPORT lean_object* l_Lean_withImportModules___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_land(size_t, size_t); +static lean_object* l___auto____x40_Lean_Environment___hyg_2443____closed__7; LEAN_EXPORT lean_object* l_Lean_Environment_header___default; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_displayStats___spec__8___lambda__1(lean_object*, lean_object*, lean_object*); lean_object* lean_save_module_data(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkEmptyEnvironment___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_setState___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2418____closed__6; lean_object* l_List_toArrayAux___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_finalizeImport___spec__11(lean_object*, size_t, size_t, lean_object*); lean_object* l___private_Lean_Data_HashMap_0__Lean_numBucketsForCapacity(lean_object*); @@ -760,6 +759,7 @@ static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__3; LEAN_EXPORT lean_object* l_Lean_ModuleIdx_toNat___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_finalizeImport___spec__8___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedModuleData___closed__2; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6214____lambda__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwAlreadyImported(lean_object*); uint8_t l_Array_isEmpty___rarg(lean_object*); static lean_object* _init_l_Lean_EnvExtensionStateSpec___closed__1() { @@ -3715,7 +3715,7 @@ x_2 = l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___closed__2; return x_2; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_initFn____x40_Lean_Environment___hyg_1289_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_initFn____x40_Lean_Environment___hyg_1314_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; @@ -3956,7 +3956,7 @@ static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg___c lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg___closed__1; x_2 = l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg___closed__2; -x_3 = lean_unsigned_to_nat(316u); +x_3 = lean_unsigned_to_nat(317u); x_4 = lean_unsigned_to_nat(4u); x_5 = l___private_Lean_Environment_0__Lean_EnvExtensionInterfaceUnsafe_invalidExtMsg; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -4021,7 +4021,7 @@ static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_modifyState___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; x_1 = l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg___closed__1; x_2 = l_Lean_EnvExtensionInterfaceUnsafe_modifyState___rarg___closed__1; -x_3 = lean_unsigned_to_nat(326u); +x_3 = lean_unsigned_to_nat(327u); x_4 = lean_unsigned_to_nat(4u); x_5 = l___private_Lean_Environment_0__Lean_EnvExtensionInterfaceUnsafe_invalidExtMsg; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -4082,7 +4082,7 @@ static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_getState___rarg___c lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg___closed__1; x_2 = l_Lean_EnvExtensionInterfaceUnsafe_getState___rarg___closed__1; -x_3 = lean_unsigned_to_nat(333u); +x_3 = lean_unsigned_to_nat(334u); x_4 = lean_unsigned_to_nat(4u); x_5 = l___private_Lean_Environment_0__Lean_EnvExtensionInterfaceUnsafe_invalidExtMsg; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -5536,7 +5536,7 @@ lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_2376_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_2401_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; @@ -5562,7 +5562,7 @@ return x_7; } } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2418____closed__1() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2443____closed__1() { _start: { lean_object* x_1; @@ -5570,7 +5570,7 @@ x_1 = lean_mk_string_from_bytes("Lean", 4); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2418____closed__2() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2443____closed__2() { _start: { lean_object* x_1; @@ -5578,7 +5578,7 @@ x_1 = lean_mk_string_from_bytes("Parser", 6); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2418____closed__3() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2443____closed__3() { _start: { lean_object* x_1; @@ -5586,7 +5586,7 @@ x_1 = lean_mk_string_from_bytes("Tactic", 6); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2418____closed__4() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2443____closed__4() { _start: { lean_object* x_1; @@ -5594,19 +5594,19 @@ x_1 = lean_mk_string_from_bytes("tacticSeq", 9); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2418____closed__5() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2443____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_Environment___hyg_2418____closed__1; -x_2 = l___auto____x40_Lean_Environment___hyg_2418____closed__2; -x_3 = l___auto____x40_Lean_Environment___hyg_2418____closed__3; -x_4 = l___auto____x40_Lean_Environment___hyg_2418____closed__4; +x_1 = l___auto____x40_Lean_Environment___hyg_2443____closed__1; +x_2 = l___auto____x40_Lean_Environment___hyg_2443____closed__2; +x_3 = l___auto____x40_Lean_Environment___hyg_2443____closed__3; +x_4 = l___auto____x40_Lean_Environment___hyg_2443____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_Environment___hyg_2418____closed__6() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2443____closed__6() { _start: { lean_object* x_1; @@ -5614,19 +5614,19 @@ x_1 = lean_mk_string_from_bytes("tacticSeq1Indented", 18); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2418____closed__7() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2443____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_Environment___hyg_2418____closed__1; -x_2 = l___auto____x40_Lean_Environment___hyg_2418____closed__2; -x_3 = l___auto____x40_Lean_Environment___hyg_2418____closed__3; -x_4 = l___auto____x40_Lean_Environment___hyg_2418____closed__6; +x_1 = l___auto____x40_Lean_Environment___hyg_2443____closed__1; +x_2 = l___auto____x40_Lean_Environment___hyg_2443____closed__2; +x_3 = l___auto____x40_Lean_Environment___hyg_2443____closed__3; +x_4 = l___auto____x40_Lean_Environment___hyg_2443____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_Environment___hyg_2418____closed__8() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2443____closed__8() { _start: { lean_object* x_1; @@ -5634,17 +5634,17 @@ x_1 = lean_mk_string_from_bytes("null", 4); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2418____closed__9() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2443____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_Environment___hyg_2418____closed__8; +x_2 = l___auto____x40_Lean_Environment___hyg_2443____closed__8; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2418____closed__10() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2443____closed__10() { _start: { lean_object* x_1; @@ -5652,41 +5652,41 @@ x_1 = lean_mk_string_from_bytes("exact", 5); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2418____closed__11() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2443____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_Environment___hyg_2418____closed__1; -x_2 = l___auto____x40_Lean_Environment___hyg_2418____closed__2; -x_3 = l___auto____x40_Lean_Environment___hyg_2418____closed__3; -x_4 = l___auto____x40_Lean_Environment___hyg_2418____closed__10; +x_1 = l___auto____x40_Lean_Environment___hyg_2443____closed__1; +x_2 = l___auto____x40_Lean_Environment___hyg_2443____closed__2; +x_3 = l___auto____x40_Lean_Environment___hyg_2443____closed__3; +x_4 = l___auto____x40_Lean_Environment___hyg_2443____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_Environment___hyg_2418____closed__12() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2443____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_Environment___hyg_2418____closed__10; +x_2 = l___auto____x40_Lean_Environment___hyg_2443____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_Environment___hyg_2418____closed__13() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2443____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_instInhabitedModuleData___closed__1; -x_2 = l___auto____x40_Lean_Environment___hyg_2418____closed__12; +x_2 = l___auto____x40_Lean_Environment___hyg_2443____closed__12; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2418____closed__14() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2443____closed__14() { _start: { lean_object* x_1; @@ -5694,7 +5694,7 @@ x_1 = lean_mk_string_from_bytes("Term", 4); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2418____closed__15() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2443____closed__15() { _start: { lean_object* x_1; @@ -5702,19 +5702,19 @@ x_1 = lean_mk_string_from_bytes("declName", 8); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2418____closed__16() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2443____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_Environment___hyg_2418____closed__1; -x_2 = l___auto____x40_Lean_Environment___hyg_2418____closed__2; -x_3 = l___auto____x40_Lean_Environment___hyg_2418____closed__14; -x_4 = l___auto____x40_Lean_Environment___hyg_2418____closed__15; +x_1 = l___auto____x40_Lean_Environment___hyg_2443____closed__1; +x_2 = l___auto____x40_Lean_Environment___hyg_2443____closed__2; +x_3 = l___auto____x40_Lean_Environment___hyg_2443____closed__14; +x_4 = l___auto____x40_Lean_Environment___hyg_2443____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_Environment___hyg_2418____closed__17() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2443____closed__17() { _start: { lean_object* x_1; @@ -5722,35 +5722,35 @@ x_1 = lean_mk_string_from_bytes("decl_name%", 10); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2418____closed__18() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2443____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_Environment___hyg_2418____closed__17; +x_2 = l___auto____x40_Lean_Environment___hyg_2443____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_Environment___hyg_2418____closed__19() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2443____closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_instInhabitedModuleData___closed__1; -x_2 = l___auto____x40_Lean_Environment___hyg_2418____closed__18; +x_2 = l___auto____x40_Lean_Environment___hyg_2443____closed__18; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2418____closed__20() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2443____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_Environment___hyg_2418____closed__16; -x_3 = l___auto____x40_Lean_Environment___hyg_2418____closed__19; +x_2 = l___auto____x40_Lean_Environment___hyg_2443____closed__16; +x_3 = l___auto____x40_Lean_Environment___hyg_2443____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); @@ -5758,23 +5758,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2418____closed__21() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2443____closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Lean_Environment___hyg_2418____closed__13; -x_2 = l___auto____x40_Lean_Environment___hyg_2418____closed__20; +x_1 = l___auto____x40_Lean_Environment___hyg_2443____closed__13; +x_2 = l___auto____x40_Lean_Environment___hyg_2443____closed__20; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2418____closed__22() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2443____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_Environment___hyg_2418____closed__11; -x_3 = l___auto____x40_Lean_Environment___hyg_2418____closed__21; +x_2 = l___auto____x40_Lean_Environment___hyg_2443____closed__11; +x_3 = l___auto____x40_Lean_Environment___hyg_2443____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); @@ -5782,23 +5782,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2418____closed__23() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2443____closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_instInhabitedModuleData___closed__1; -x_2 = l___auto____x40_Lean_Environment___hyg_2418____closed__22; +x_2 = l___auto____x40_Lean_Environment___hyg_2443____closed__22; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2418____closed__24() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2443____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_Environment___hyg_2418____closed__9; -x_3 = l___auto____x40_Lean_Environment___hyg_2418____closed__23; +x_2 = l___auto____x40_Lean_Environment___hyg_2443____closed__9; +x_3 = l___auto____x40_Lean_Environment___hyg_2443____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); @@ -5806,23 +5806,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2418____closed__25() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2443____closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_instInhabitedModuleData___closed__1; -x_2 = l___auto____x40_Lean_Environment___hyg_2418____closed__24; +x_2 = l___auto____x40_Lean_Environment___hyg_2443____closed__24; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2418____closed__26() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2443____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_Environment___hyg_2418____closed__7; -x_3 = l___auto____x40_Lean_Environment___hyg_2418____closed__25; +x_2 = l___auto____x40_Lean_Environment___hyg_2443____closed__7; +x_3 = l___auto____x40_Lean_Environment___hyg_2443____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); @@ -5830,23 +5830,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2418____closed__27() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2443____closed__27() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_instInhabitedModuleData___closed__1; -x_2 = l___auto____x40_Lean_Environment___hyg_2418____closed__26; +x_2 = l___auto____x40_Lean_Environment___hyg_2443____closed__26; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2418____closed__28() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2443____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_Environment___hyg_2418____closed__5; -x_3 = l___auto____x40_Lean_Environment___hyg_2418____closed__27; +x_2 = l___auto____x40_Lean_Environment___hyg_2443____closed__5; +x_3 = l___auto____x40_Lean_Environment___hyg_2443____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); @@ -5854,11 +5854,11 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2418_() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2443_() { _start: { lean_object* x_1; -x_1 = l___auto____x40_Lean_Environment___hyg_2418____closed__28; +x_1 = l___auto____x40_Lean_Environment___hyg_2443____closed__28; return x_1; } } @@ -6316,11 +6316,11 @@ x_3 = lean_alloc_closure((void*)(l_Lean_mkStateFromImportedEntries___rarg), 3, 0 return x_3; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2821_() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2846_() { _start: { lean_object* x_1; -x_1 = l___auto____x40_Lean_Environment___hyg_2418____closed__28; +x_1 = l___auto____x40_Lean_Environment___hyg_2443____closed__28; return x_1; } } @@ -6728,11 +6728,11 @@ lean_dec(x_1); return x_4; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3160_() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3185_() { _start: { lean_object* x_1; -x_1 = l___auto____x40_Lean_Environment___hyg_2418____closed__28; +x_1 = l___auto____x40_Lean_Environment___hyg_2443____closed__28; return x_1; } } @@ -6944,7 +6944,7 @@ static lean_object* _init_l_Lean_TagDeclarationExtension_tag___closed__5() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg___closed__1; x_2 = l_Lean_TagDeclarationExtension_tag___closed__4; -x_3 = lean_unsigned_to_nat(581u); +x_3 = lean_unsigned_to_nat(582u); x_4 = lean_unsigned_to_nat(2u); x_5 = l_Lean_TagDeclarationExtension_tag___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -7143,11 +7143,11 @@ x_5 = lean_box(x_4); return x_5; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3338_() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3363_() { _start: { lean_object* x_1; -x_1 = l___auto____x40_Lean_Environment___hyg_2418____closed__28; +x_1 = l___auto____x40_Lean_Environment___hyg_2443____closed__28; return x_1; } } @@ -7401,7 +7401,7 @@ static lean_object* _init_l_Lean_MapDeclarationExtension_insert___rarg___closed_ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg___closed__1; x_2 = l_Lean_MapDeclarationExtension_insert___rarg___closed__3; -x_3 = lean_unsigned_to_nat(611u); +x_3 = lean_unsigned_to_nat(612u); x_4 = lean_unsigned_to_nat(2u); x_5 = l_Lean_MapDeclarationExtension_insert___rarg___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -13549,7 +13549,7 @@ x_7 = l_Lean_withImportModules___rarg(x_1, x_2, x_6, x_4, x_5); return x_7; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6189____spec__2(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__2(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { uint8_t x_5; @@ -13572,7 +13572,7 @@ return x_4; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6189____spec__3(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_initFn____x40_Lean_Environment___hyg_6214____spec__3(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { uint8_t x_5; @@ -13610,7 +13610,7 @@ size_t x_15; size_t x_16; lean_object* x_17; x_15 = 0; x_16 = lean_usize_of_nat(x_7); lean_dec(x_7); -x_17 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6189____spec__2(x_6, x_15, x_16, x_4); +x_17 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__2(x_6, x_15, x_16, x_4); lean_dec(x_6); x_2 = x_11; x_4 = x_17; @@ -13624,7 +13624,7 @@ return x_4; } } } -LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6189____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; uint8_t x_5; @@ -13653,14 +13653,14 @@ size_t x_7; size_t x_8; lean_object* x_9; x_7 = 0; x_8 = lean_usize_of_nat(x_3); lean_dec(x_3); -x_9 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6189____spec__3(x_2, x_7, x_8, x_1); +x_9 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__3(x_2, x_7, x_8, x_1); lean_dec(x_2); return x_9; } } } } -LEAN_EXPORT lean_object* l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6189____spec__4(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__4(lean_object* x_1) { _start: { uint8_t x_2; @@ -13698,7 +13698,7 @@ return x_8; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6189____spec__5(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_initFn____x40_Lean_Environment___hyg_6214____spec__5(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { uint8_t x_5; @@ -13724,7 +13724,7 @@ return x_4; } } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6189____lambda__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6214____lambda__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; @@ -13733,7 +13733,7 @@ x_4 = l_Lean_SMap_insert___at_Lean_NameSSet_insert___spec__1(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6189____lambda__2___closed__1() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6214____lambda__2___closed__1() { _start: { lean_object* x_1; lean_object* x_2; @@ -13742,7 +13742,7 @@ x_2 = l_Lean_mkHashMapImp___rarg(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6189____lambda__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6214____lambda__2(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; @@ -13753,15 +13753,15 @@ if (x_4 == 0) { 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_dec(x_2); -x_5 = l_Lean_initFn____x40_Lean_Environment___hyg_6189____lambda__2___closed__1; -x_6 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6189____spec__1(x_5, x_1); +x_5 = l_Lean_initFn____x40_Lean_Environment___hyg_6214____lambda__2___closed__1; +x_6 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__1(x_5, x_1); x_7 = 1; x_8 = l_Lean_mkEmptyEnvironment___lambda__1___closed__3; x_9 = lean_alloc_ctor(0, 2, 1); lean_ctor_set(x_9, 0, x_6); lean_ctor_set(x_9, 1, x_8); lean_ctor_set_uint8(x_9, sizeof(void*)*2, x_7); -x_10 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6189____spec__4(x_9); +x_10 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__4(x_9); return x_10; } else @@ -13772,15 +13772,15 @@ if (x_11 == 0) { lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_dec(x_2); -x_12 = l_Lean_initFn____x40_Lean_Environment___hyg_6189____lambda__2___closed__1; -x_13 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6189____spec__1(x_12, x_1); +x_12 = l_Lean_initFn____x40_Lean_Environment___hyg_6214____lambda__2___closed__1; +x_13 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__1(x_12, x_1); x_14 = 1; x_15 = l_Lean_mkEmptyEnvironment___lambda__1___closed__3; x_16 = lean_alloc_ctor(0, 2, 1); lean_ctor_set(x_16, 0, x_13); lean_ctor_set(x_16, 1, x_15); lean_ctor_set_uint8(x_16, sizeof(void*)*2, x_14); -x_17 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6189____spec__4(x_16); +x_17 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__4(x_16); return x_17; } else @@ -13789,23 +13789,23 @@ size_t x_18; size_t x_19; lean_object* x_20; lean_object* x_21; lean_object* x_2 x_18 = 0; x_19 = lean_usize_of_nat(x_2); lean_dec(x_2); -x_20 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6189____spec__5(x_1, x_18, x_19, x_3); +x_20 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__5(x_1, x_18, x_19, x_3); x_21 = l_Lean_mkHashMapImp___rarg(x_20); lean_dec(x_20); -x_22 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6189____spec__1(x_21, x_1); +x_22 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__1(x_21, x_1); x_23 = 1; x_24 = l_Lean_mkEmptyEnvironment___lambda__1___closed__3; x_25 = lean_alloc_ctor(0, 2, 1); lean_ctor_set(x_25, 0, x_22); lean_ctor_set(x_25, 1, x_24); lean_ctor_set_uint8(x_25, sizeof(void*)*2, x_23); -x_26 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6189____spec__4(x_25); +x_26 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__4(x_25); return x_26; } } } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6189____closed__1() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__1() { _start: { lean_object* x_1; @@ -13813,17 +13813,17 @@ x_1 = lean_mk_string_from_bytes("namespacesExt", 13); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6189____closed__2() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Lean_Environment___hyg_2418____closed__1; -x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_6189____closed__1; +x_1 = l___auto____x40_Lean_Environment___hyg_2443____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6189____closed__3() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__3() { _start: { lean_object* x_1; @@ -13832,30 +13832,30 @@ lean_closure_set(x_1, 0, lean_box(0)); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6189____closed__4() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__4() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_initFn____x40_Lean_Environment___hyg_6189____lambda__1), 2, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_initFn____x40_Lean_Environment___hyg_6214____lambda__1), 2, 0); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6189____closed__5() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__5() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_initFn____x40_Lean_Environment___hyg_6189____lambda__2), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_initFn____x40_Lean_Environment___hyg_6214____lambda__2), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6189____closed__6() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6214____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_initFn____x40_Lean_Environment___hyg_6189____closed__2; -x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_6189____closed__4; -x_3 = l_Lean_initFn____x40_Lean_Environment___hyg_6189____closed__5; -x_4 = l_Lean_initFn____x40_Lean_Environment___hyg_6189____closed__3; +x_1 = l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__2; +x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__4; +x_3 = l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__5; +x_4 = l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__3; x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -13864,16 +13864,16 @@ lean_ctor_set(x_5, 3, x_4); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6189_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6214_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_6189____closed__6; +x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__6; x_3 = l_Lean_registerSimplePersistentEnvExtension___rarg(x_2, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6189____spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__2___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; @@ -13881,12 +13881,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_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6189____spec__2(x_1, x_5, x_6, x_4); +x_7 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__2(x_1, x_5, x_6, x_4); lean_dec(x_1); return x_7; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6189____spec__3___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_initFn____x40_Lean_Environment___hyg_6214____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; @@ -13894,12 +13894,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_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6189____spec__3(x_1, x_5, x_6, x_4); +x_7 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__3(x_1, x_5, x_6, x_4); lean_dec(x_1); return x_7; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6189____spec__5___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_initFn____x40_Lean_Environment___hyg_6214____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; @@ -13907,7 +13907,7 @@ x_5 = lean_unbox_usize(x_2); lean_dec(x_2); x_6 = lean_unbox_usize(x_3); lean_dec(x_3); -x_7 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6189____spec__5(x_1, x_5, x_6, x_4); +x_7 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__5(x_1, x_5, x_6, x_4); lean_dec(x_1); return x_7; } @@ -15835,7 +15835,7 @@ l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___closed__1 = _init_l_Lean_E lean_mark_persistent(l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___closed__1); l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___closed__2 = _init_l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___closed__2(); lean_mark_persistent(l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___closed__2); -if (builtin) {res = l_Lean_EnvExtensionInterfaceUnsafe_initFn____x40_Lean_Environment___hyg_1289_(lean_io_mk_world()); +if (builtin) {res = l_Lean_EnvExtensionInterfaceUnsafe_initFn____x40_Lean_Environment___hyg_1314_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l___private_Lean_Environment_0__Lean_EnvExtensionInterfaceUnsafe_envExtensionsRef = lean_io_result_get_value(res); lean_mark_persistent(l___private_Lean_Environment_0__Lean_EnvExtensionInterfaceUnsafe_envExtensionsRef); @@ -15902,69 +15902,69 @@ l_Lean_instInhabitedPersistentEnvExtension___closed__4 = _init_l_Lean_instInhabi lean_mark_persistent(l_Lean_instInhabitedPersistentEnvExtension___closed__4); l_Lean_instInhabitedPersistentEnvExtension___closed__5 = _init_l_Lean_instInhabitedPersistentEnvExtension___closed__5(); lean_mark_persistent(l_Lean_instInhabitedPersistentEnvExtension___closed__5); -if (builtin) {res = l_Lean_initFn____x40_Lean_Environment___hyg_2376_(lean_io_mk_world()); +if (builtin) {res = l_Lean_initFn____x40_Lean_Environment___hyg_2401_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_persistentEnvExtensionsRef = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_persistentEnvExtensionsRef); lean_dec_ref(res); -}l___auto____x40_Lean_Environment___hyg_2418____closed__1 = _init_l___auto____x40_Lean_Environment___hyg_2418____closed__1(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2418____closed__1); -l___auto____x40_Lean_Environment___hyg_2418____closed__2 = _init_l___auto____x40_Lean_Environment___hyg_2418____closed__2(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2418____closed__2); -l___auto____x40_Lean_Environment___hyg_2418____closed__3 = _init_l___auto____x40_Lean_Environment___hyg_2418____closed__3(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2418____closed__3); -l___auto____x40_Lean_Environment___hyg_2418____closed__4 = _init_l___auto____x40_Lean_Environment___hyg_2418____closed__4(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2418____closed__4); -l___auto____x40_Lean_Environment___hyg_2418____closed__5 = _init_l___auto____x40_Lean_Environment___hyg_2418____closed__5(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2418____closed__5); -l___auto____x40_Lean_Environment___hyg_2418____closed__6 = _init_l___auto____x40_Lean_Environment___hyg_2418____closed__6(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2418____closed__6); -l___auto____x40_Lean_Environment___hyg_2418____closed__7 = _init_l___auto____x40_Lean_Environment___hyg_2418____closed__7(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2418____closed__7); -l___auto____x40_Lean_Environment___hyg_2418____closed__8 = _init_l___auto____x40_Lean_Environment___hyg_2418____closed__8(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2418____closed__8); -l___auto____x40_Lean_Environment___hyg_2418____closed__9 = _init_l___auto____x40_Lean_Environment___hyg_2418____closed__9(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2418____closed__9); -l___auto____x40_Lean_Environment___hyg_2418____closed__10 = _init_l___auto____x40_Lean_Environment___hyg_2418____closed__10(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2418____closed__10); -l___auto____x40_Lean_Environment___hyg_2418____closed__11 = _init_l___auto____x40_Lean_Environment___hyg_2418____closed__11(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2418____closed__11); -l___auto____x40_Lean_Environment___hyg_2418____closed__12 = _init_l___auto____x40_Lean_Environment___hyg_2418____closed__12(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2418____closed__12); -l___auto____x40_Lean_Environment___hyg_2418____closed__13 = _init_l___auto____x40_Lean_Environment___hyg_2418____closed__13(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2418____closed__13); -l___auto____x40_Lean_Environment___hyg_2418____closed__14 = _init_l___auto____x40_Lean_Environment___hyg_2418____closed__14(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2418____closed__14); -l___auto____x40_Lean_Environment___hyg_2418____closed__15 = _init_l___auto____x40_Lean_Environment___hyg_2418____closed__15(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2418____closed__15); -l___auto____x40_Lean_Environment___hyg_2418____closed__16 = _init_l___auto____x40_Lean_Environment___hyg_2418____closed__16(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2418____closed__16); -l___auto____x40_Lean_Environment___hyg_2418____closed__17 = _init_l___auto____x40_Lean_Environment___hyg_2418____closed__17(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2418____closed__17); -l___auto____x40_Lean_Environment___hyg_2418____closed__18 = _init_l___auto____x40_Lean_Environment___hyg_2418____closed__18(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2418____closed__18); -l___auto____x40_Lean_Environment___hyg_2418____closed__19 = _init_l___auto____x40_Lean_Environment___hyg_2418____closed__19(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2418____closed__19); -l___auto____x40_Lean_Environment___hyg_2418____closed__20 = _init_l___auto____x40_Lean_Environment___hyg_2418____closed__20(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2418____closed__20); -l___auto____x40_Lean_Environment___hyg_2418____closed__21 = _init_l___auto____x40_Lean_Environment___hyg_2418____closed__21(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2418____closed__21); -l___auto____x40_Lean_Environment___hyg_2418____closed__22 = _init_l___auto____x40_Lean_Environment___hyg_2418____closed__22(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2418____closed__22); -l___auto____x40_Lean_Environment___hyg_2418____closed__23 = _init_l___auto____x40_Lean_Environment___hyg_2418____closed__23(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2418____closed__23); -l___auto____x40_Lean_Environment___hyg_2418____closed__24 = _init_l___auto____x40_Lean_Environment___hyg_2418____closed__24(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2418____closed__24); -l___auto____x40_Lean_Environment___hyg_2418____closed__25 = _init_l___auto____x40_Lean_Environment___hyg_2418____closed__25(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2418____closed__25); -l___auto____x40_Lean_Environment___hyg_2418____closed__26 = _init_l___auto____x40_Lean_Environment___hyg_2418____closed__26(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2418____closed__26); -l___auto____x40_Lean_Environment___hyg_2418____closed__27 = _init_l___auto____x40_Lean_Environment___hyg_2418____closed__27(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2418____closed__27); -l___auto____x40_Lean_Environment___hyg_2418____closed__28 = _init_l___auto____x40_Lean_Environment___hyg_2418____closed__28(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2418____closed__28); -l___auto____x40_Lean_Environment___hyg_2418_ = _init_l___auto____x40_Lean_Environment___hyg_2418_(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2418_); +}l___auto____x40_Lean_Environment___hyg_2443____closed__1 = _init_l___auto____x40_Lean_Environment___hyg_2443____closed__1(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2443____closed__1); +l___auto____x40_Lean_Environment___hyg_2443____closed__2 = _init_l___auto____x40_Lean_Environment___hyg_2443____closed__2(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2443____closed__2); +l___auto____x40_Lean_Environment___hyg_2443____closed__3 = _init_l___auto____x40_Lean_Environment___hyg_2443____closed__3(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2443____closed__3); +l___auto____x40_Lean_Environment___hyg_2443____closed__4 = _init_l___auto____x40_Lean_Environment___hyg_2443____closed__4(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2443____closed__4); +l___auto____x40_Lean_Environment___hyg_2443____closed__5 = _init_l___auto____x40_Lean_Environment___hyg_2443____closed__5(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2443____closed__5); +l___auto____x40_Lean_Environment___hyg_2443____closed__6 = _init_l___auto____x40_Lean_Environment___hyg_2443____closed__6(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2443____closed__6); +l___auto____x40_Lean_Environment___hyg_2443____closed__7 = _init_l___auto____x40_Lean_Environment___hyg_2443____closed__7(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2443____closed__7); +l___auto____x40_Lean_Environment___hyg_2443____closed__8 = _init_l___auto____x40_Lean_Environment___hyg_2443____closed__8(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2443____closed__8); +l___auto____x40_Lean_Environment___hyg_2443____closed__9 = _init_l___auto____x40_Lean_Environment___hyg_2443____closed__9(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2443____closed__9); +l___auto____x40_Lean_Environment___hyg_2443____closed__10 = _init_l___auto____x40_Lean_Environment___hyg_2443____closed__10(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2443____closed__10); +l___auto____x40_Lean_Environment___hyg_2443____closed__11 = _init_l___auto____x40_Lean_Environment___hyg_2443____closed__11(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2443____closed__11); +l___auto____x40_Lean_Environment___hyg_2443____closed__12 = _init_l___auto____x40_Lean_Environment___hyg_2443____closed__12(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2443____closed__12); +l___auto____x40_Lean_Environment___hyg_2443____closed__13 = _init_l___auto____x40_Lean_Environment___hyg_2443____closed__13(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2443____closed__13); +l___auto____x40_Lean_Environment___hyg_2443____closed__14 = _init_l___auto____x40_Lean_Environment___hyg_2443____closed__14(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2443____closed__14); +l___auto____x40_Lean_Environment___hyg_2443____closed__15 = _init_l___auto____x40_Lean_Environment___hyg_2443____closed__15(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2443____closed__15); +l___auto____x40_Lean_Environment___hyg_2443____closed__16 = _init_l___auto____x40_Lean_Environment___hyg_2443____closed__16(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2443____closed__16); +l___auto____x40_Lean_Environment___hyg_2443____closed__17 = _init_l___auto____x40_Lean_Environment___hyg_2443____closed__17(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2443____closed__17); +l___auto____x40_Lean_Environment___hyg_2443____closed__18 = _init_l___auto____x40_Lean_Environment___hyg_2443____closed__18(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2443____closed__18); +l___auto____x40_Lean_Environment___hyg_2443____closed__19 = _init_l___auto____x40_Lean_Environment___hyg_2443____closed__19(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2443____closed__19); +l___auto____x40_Lean_Environment___hyg_2443____closed__20 = _init_l___auto____x40_Lean_Environment___hyg_2443____closed__20(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2443____closed__20); +l___auto____x40_Lean_Environment___hyg_2443____closed__21 = _init_l___auto____x40_Lean_Environment___hyg_2443____closed__21(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2443____closed__21); +l___auto____x40_Lean_Environment___hyg_2443____closed__22 = _init_l___auto____x40_Lean_Environment___hyg_2443____closed__22(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2443____closed__22); +l___auto____x40_Lean_Environment___hyg_2443____closed__23 = _init_l___auto____x40_Lean_Environment___hyg_2443____closed__23(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2443____closed__23); +l___auto____x40_Lean_Environment___hyg_2443____closed__24 = _init_l___auto____x40_Lean_Environment___hyg_2443____closed__24(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2443____closed__24); +l___auto____x40_Lean_Environment___hyg_2443____closed__25 = _init_l___auto____x40_Lean_Environment___hyg_2443____closed__25(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2443____closed__25); +l___auto____x40_Lean_Environment___hyg_2443____closed__26 = _init_l___auto____x40_Lean_Environment___hyg_2443____closed__26(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2443____closed__26); +l___auto____x40_Lean_Environment___hyg_2443____closed__27 = _init_l___auto____x40_Lean_Environment___hyg_2443____closed__27(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2443____closed__27); +l___auto____x40_Lean_Environment___hyg_2443____closed__28 = _init_l___auto____x40_Lean_Environment___hyg_2443____closed__28(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2443____closed__28); +l___auto____x40_Lean_Environment___hyg_2443_ = _init_l___auto____x40_Lean_Environment___hyg_2443_(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2443_); l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___closed__1 = _init_l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___closed__1(); lean_mark_persistent(l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___closed__1); l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___closed__2 = _init_l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___closed__2(); @@ -15973,16 +15973,16 @@ l_Lean_registerPersistentEnvExtensionUnsafe___rarg___closed__1 = _init_l_Lean_re lean_mark_persistent(l_Lean_registerPersistentEnvExtensionUnsafe___rarg___closed__1); l_Lean_registerPersistentEnvExtensionUnsafe___rarg___closed__2 = _init_l_Lean_registerPersistentEnvExtensionUnsafe___rarg___closed__2(); lean_mark_persistent(l_Lean_registerPersistentEnvExtensionUnsafe___rarg___closed__2); -l___auto____x40_Lean_Environment___hyg_2821_ = _init_l___auto____x40_Lean_Environment___hyg_2821_(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2821_); +l___auto____x40_Lean_Environment___hyg_2846_ = _init_l___auto____x40_Lean_Environment___hyg_2846_(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2846_); l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4___closed__1 = _init_l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4___closed__1(); lean_mark_persistent(l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4___closed__1); l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4___closed__2 = _init_l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4___closed__2(); lean_mark_persistent(l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4___closed__2); l_Lean_registerSimplePersistentEnvExtension___rarg___closed__1 = _init_l_Lean_registerSimplePersistentEnvExtension___rarg___closed__1(); lean_mark_persistent(l_Lean_registerSimplePersistentEnvExtension___rarg___closed__1); -l___auto____x40_Lean_Environment___hyg_3160_ = _init_l___auto____x40_Lean_Environment___hyg_3160_(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3160_); +l___auto____x40_Lean_Environment___hyg_3185_ = _init_l___auto____x40_Lean_Environment___hyg_3185_(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3185_); l_Array_qsort_sort___at_Lean_mkTagDeclarationExtension___spec__1___closed__1 = _init_l_Array_qsort_sort___at_Lean_mkTagDeclarationExtension___spec__1___closed__1(); lean_mark_persistent(l_Array_qsort_sort___at_Lean_mkTagDeclarationExtension___spec__1___closed__1); l_Lean_mkTagDeclarationExtension___closed__1 = _init_l_Lean_mkTagDeclarationExtension___closed__1(); @@ -16007,8 +16007,8 @@ l_Lean_TagDeclarationExtension_tag___closed__5 = _init_l_Lean_TagDeclarationExte lean_mark_persistent(l_Lean_TagDeclarationExtension_tag___closed__5); l_Lean_TagDeclarationExtension_isTagged___closed__1 = _init_l_Lean_TagDeclarationExtension_isTagged___closed__1(); lean_mark_persistent(l_Lean_TagDeclarationExtension_isTagged___closed__1); -l___auto____x40_Lean_Environment___hyg_3338_ = _init_l___auto____x40_Lean_Environment___hyg_3338_(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3338_); +l___auto____x40_Lean_Environment___hyg_3363_ = _init_l___auto____x40_Lean_Environment___hyg_3363_(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3363_); l_Array_qsort_sort___at_Lean_mkMapDeclarationExtension___spec__1___rarg___closed__1 = _init_l_Array_qsort_sort___at_Lean_mkMapDeclarationExtension___spec__1___rarg___closed__1(); lean_mark_persistent(l_Array_qsort_sort___at_Lean_mkMapDeclarationExtension___spec__1___rarg___closed__1); l_Lean_mkMapDeclarationExtension___rarg___closed__1 = _init_l_Lean_mkMapDeclarationExtension___rarg___closed__1(); @@ -16081,21 +16081,21 @@ l_Lean_importModules___closed__1 = _init_l_Lean_importModules___closed__1(); lean_mark_persistent(l_Lean_importModules___closed__1); l_Lean_importModules___boxed__const__1 = _init_l_Lean_importModules___boxed__const__1(); lean_mark_persistent(l_Lean_importModules___boxed__const__1); -l_Lean_initFn____x40_Lean_Environment___hyg_6189____lambda__2___closed__1 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6189____lambda__2___closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6189____lambda__2___closed__1); -l_Lean_initFn____x40_Lean_Environment___hyg_6189____closed__1 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6189____closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6189____closed__1); -l_Lean_initFn____x40_Lean_Environment___hyg_6189____closed__2 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6189____closed__2(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6189____closed__2); -l_Lean_initFn____x40_Lean_Environment___hyg_6189____closed__3 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6189____closed__3(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6189____closed__3); -l_Lean_initFn____x40_Lean_Environment___hyg_6189____closed__4 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6189____closed__4(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6189____closed__4); -l_Lean_initFn____x40_Lean_Environment___hyg_6189____closed__5 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6189____closed__5(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6189____closed__5); -l_Lean_initFn____x40_Lean_Environment___hyg_6189____closed__6 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6189____closed__6(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6189____closed__6); -if (builtin) {res = l_Lean_initFn____x40_Lean_Environment___hyg_6189_(lean_io_mk_world()); +l_Lean_initFn____x40_Lean_Environment___hyg_6214____lambda__2___closed__1 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6214____lambda__2___closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6214____lambda__2___closed__1); +l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__1 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__1); +l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__2 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__2); +l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__3 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__3); +l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__4 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__4(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__4); +l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__5 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__5(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__5); +l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__6 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__6(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__6); +if (builtin) {res = l_Lean_initFn____x40_Lean_Environment___hyg_6214_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_namespacesExt = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_namespacesExt); diff --git a/stage0/stdlib/Lean/Linter/MissingDocs.c b/stage0/stdlib/Lean/Linter/MissingDocs.c index 6b51e3049f69..3ca6af706499 100644 --- a/stage0/stdlib/Lean/Linter/MissingDocs.c +++ b/stage0/stdlib/Lean/Linter/MissingDocs.c @@ -174,6 +174,7 @@ 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_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_object* l_instDecidableEqPos___boxed(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*); @@ -369,7 +370,6 @@ static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingD 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_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_object* l_List_reverse___rarg(lean_object*); @@ -5603,7 +5603,7 @@ static lean_object* _init_l_Lean_Linter_MissingDocs_checkDecl___lambda__1___clos _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Nat_decEq___boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l_instDecidableEqPos___boxed), 2, 0); return x_1; } } diff --git a/stage0/stdlib/Lean/Message.c b/stage0/stdlib/Lean/Message.c index 85724535ad28..a079f66093b1 100644 --- a/stage0/stdlib/Lean/Message.c +++ b/stage0/stdlib/Lean/Message.c @@ -223,6 +223,7 @@ static lean_object* l_Lean_MessageData_instCoeArrayExprMessageData___closed__3; uint8_t l_Std_Format_isEmpty(lean_object*); static lean_object* l_Lean_MessageData_instCoeArrayExprMessageData___closed__2; LEAN_EXPORT lean_object* l_String_split___at_Lean_stringToMessageData___spec__1___boxed(lean_object*); +static lean_object* l_Lean_KernelException_toMessageData___closed__53; LEAN_EXPORT lean_object* l_Lean_instToMessageDataProd___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addMessageContextFull___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_MessageData_ofSyntax___lambda__2(lean_object*); @@ -404,6 +405,7 @@ LEAN_EXPORT lean_object* l_Lean_MessageData_arrayExpr_toMessageData___boxed(lean LEAN_EXPORT lean_object* l_Lean_KernelException_toMessageData(lean_object*, lean_object*); static lean_object* l_Lean_MessageData_hasSyntheticSorry_visit___closed__5; static lean_object* l_Lean_MessageData_instCoeArrayExprMessageData___closed__1; +static lean_object* l_Lean_KernelException_toMessageData___closed__51; lean_object* l_String_toSubstring_x27(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_MessageLog_errorsToWarnings___spec__3(size_t, size_t, lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); @@ -411,6 +413,7 @@ static lean_object* l_Lean_instInhabitedMessageLog___closed__1; LEAN_EXPORT lean_object* l_Lean_MessageData_ofName(lean_object*); lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MessageData_paren___closed__2; +static lean_object* l_Lean_KernelException_toMessageData___closed__52; static lean_object* l_Lean_instInhabitedMessage___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_MessageLog_errorsToWarnings___spec__4___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Nat_repr(lean_object*); @@ -423,6 +426,7 @@ static lean_object* l_Lean___aux__Lean__Message______macroRules__Lean__termM_x21 static lean_object* l_Lean_KernelException_toMessageData___closed__27; LEAN_EXPORT lean_object* l_Lean_MessageData_instCoeNameMessageData; LEAN_EXPORT lean_object* l_Lean_stringToMessageData___boxed(lean_object*); +static lean_object* l_Lean_KernelException_toMessageData___closed__54; LEAN_EXPORT lean_object* l_Lean_MessageData_instCoeFormatMessageData(lean_object*); LEAN_EXPORT lean_object* l_Lean_addMessageContextFull___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Message_toString(lean_object*, uint8_t, lean_object*); @@ -7096,7 +7100,7 @@ static lean_object* _init_l_Lean_KernelException_toMessageData___closed__37() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("(kernel) ", 9); +x_1 = lean_mk_string_from_bytes("(kernel) type of theorem '", 26); return x_1; } } @@ -7113,7 +7117,7 @@ static lean_object* _init_l_Lean_KernelException_toMessageData___closed__39() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("(kernel) deterministic timeout", 30); +x_1 = lean_mk_string_from_bytes("' is not a proposition", 22); return x_1; } } @@ -7122,22 +7126,56 @@ static lean_object* _init_l_Lean_KernelException_toMessageData___closed__40() { { lean_object* x_1; lean_object* x_2; x_1 = l_Lean_KernelException_toMessageData___closed__39; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_KernelException_toMessageData___closed__41() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("(kernel) ", 9); +return x_1; +} +} +static lean_object* _init_l_Lean_KernelException_toMessageData___closed__42() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_KernelException_toMessageData___closed__41; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_KernelException_toMessageData___closed__43() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("(kernel) deterministic timeout", 30); +return x_1; +} +} +static lean_object* _init_l_Lean_KernelException_toMessageData___closed__44() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_KernelException_toMessageData___closed__43; 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_KernelException_toMessageData___closed__41() { +static lean_object* _init_l_Lean_KernelException_toMessageData___closed__45() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_KernelException_toMessageData___closed__40; +x_1 = l_Lean_KernelException_toMessageData___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_KernelException_toMessageData___closed__42() { +static lean_object* _init_l_Lean_KernelException_toMessageData___closed__46() { _start: { lean_object* x_1; @@ -7145,27 +7183,27 @@ x_1 = lean_mk_string_from_bytes("(kernel) excessive memory consumption detected" return x_1; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__43() { +static lean_object* _init_l_Lean_KernelException_toMessageData___closed__47() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_KernelException_toMessageData___closed__42; +x_1 = l_Lean_KernelException_toMessageData___closed__46; 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_KernelException_toMessageData___closed__44() { +static lean_object* _init_l_Lean_KernelException_toMessageData___closed__48() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_KernelException_toMessageData___closed__43; +x_1 = l_Lean_KernelException_toMessageData___closed__47; 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_KernelException_toMessageData___closed__45() { +static lean_object* _init_l_Lean_KernelException_toMessageData___closed__49() { _start: { lean_object* x_1; @@ -7173,27 +7211,27 @@ x_1 = lean_mk_string_from_bytes("(kernel) deep recursion detected", 32); return x_1; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__46() { +static lean_object* _init_l_Lean_KernelException_toMessageData___closed__50() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_KernelException_toMessageData___closed__45; +x_1 = l_Lean_KernelException_toMessageData___closed__49; 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_KernelException_toMessageData___closed__47() { +static lean_object* _init_l_Lean_KernelException_toMessageData___closed__51() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_KernelException_toMessageData___closed__46; +x_1 = l_Lean_KernelException_toMessageData___closed__50; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__48() { +static lean_object* _init_l_Lean_KernelException_toMessageData___closed__52() { _start: { lean_object* x_1; @@ -7201,21 +7239,21 @@ x_1 = lean_mk_string_from_bytes("(kernel) interrupted", 20); return x_1; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__49() { +static lean_object* _init_l_Lean_KernelException_toMessageData___closed__53() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_KernelException_toMessageData___closed__48; +x_1 = l_Lean_KernelException_toMessageData___closed__52; 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_KernelException_toMessageData___closed__50() { +static lean_object* _init_l_Lean_KernelException_toMessageData___closed__54() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_KernelException_toMessageData___closed__49; +x_1 = l_Lean_KernelException_toMessageData___closed__53; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; @@ -7579,50 +7617,81 @@ return x_149; } case 11: { -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_dec(x_2); +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_150 = lean_ctor_get(x_1, 0); lean_inc(x_150); +x_151 = lean_ctor_get(x_1, 1); +lean_inc(x_151); +x_152 = lean_ctor_get(x_1, 2); +lean_inc(x_152); lean_dec(x_1); -x_151 = l_Lean_stringToMessageData(x_150); -lean_dec(x_150); -x_152 = l_Lean_KernelException_toMessageData___closed__38; -x_153 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_153, 0, x_152); -lean_ctor_set(x_153, 1, x_151); -x_154 = l_Lean_KernelException_toMessageData___closed__16; +x_153 = l_Lean_MessageData_ofName(x_151); +x_154 = l_Lean_KernelException_toMessageData___closed__38; x_155 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_155, 0, x_153); -lean_ctor_set(x_155, 1, x_154); -return x_155; +lean_ctor_set(x_155, 0, x_154); +lean_ctor_set(x_155, 1, x_153); +x_156 = l_Lean_KernelException_toMessageData___closed__40; +x_157 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_157, 0, x_155); +lean_ctor_set(x_157, 1, x_156); +x_158 = l_Lean_indentExpr(x_152); +x_159 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_159, 0, x_157); +lean_ctor_set(x_159, 1, x_158); +x_160 = l_Lean_KernelException_toMessageData___closed__16; +x_161 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_161, 0, x_159); +lean_ctor_set(x_161, 1, x_160); +x_162 = l_Lean_addMessageContextPartial___rarg___lambda__1___closed__2; +x_163 = l___private_Lean_Message_0__Lean_KernelException_mkCtx(x_150, x_162, x_2, x_161); +return x_163; } case 12: { -lean_object* x_156; +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_dec(x_2); -x_156 = l_Lean_KernelException_toMessageData___closed__41; -return x_156; +x_164 = lean_ctor_get(x_1, 0); +lean_inc(x_164); +lean_dec(x_1); +x_165 = l_Lean_stringToMessageData(x_164); +lean_dec(x_164); +x_166 = l_Lean_KernelException_toMessageData___closed__42; +x_167 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_167, 0, x_166); +lean_ctor_set(x_167, 1, x_165); +x_168 = l_Lean_KernelException_toMessageData___closed__16; +x_169 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_169, 0, x_167); +lean_ctor_set(x_169, 1, x_168); +return x_169; } case 13: { -lean_object* x_157; +lean_object* x_170; lean_dec(x_2); -x_157 = l_Lean_KernelException_toMessageData___closed__44; -return x_157; +x_170 = l_Lean_KernelException_toMessageData___closed__45; +return x_170; } case 14: { -lean_object* x_158; +lean_object* x_171; lean_dec(x_2); -x_158 = l_Lean_KernelException_toMessageData___closed__47; -return x_158; +x_171 = l_Lean_KernelException_toMessageData___closed__48; +return x_171; +} +case 15: +{ +lean_object* x_172; +lean_dec(x_2); +x_172 = l_Lean_KernelException_toMessageData___closed__51; +return x_172; } default: { -lean_object* x_159; +lean_object* x_173; lean_dec(x_2); -x_159 = l_Lean_KernelException_toMessageData___closed__50; -return x_159; +x_173 = l_Lean_KernelException_toMessageData___closed__54; +return x_173; } } } @@ -8034,6 +8103,14 @@ l_Lean_KernelException_toMessageData___closed__49 = _init_l_Lean_KernelException lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__49); l_Lean_KernelException_toMessageData___closed__50 = _init_l_Lean_KernelException_toMessageData___closed__50(); lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__50); +l_Lean_KernelException_toMessageData___closed__51 = _init_l_Lean_KernelException_toMessageData___closed__51(); +lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__51); +l_Lean_KernelException_toMessageData___closed__52 = _init_l_Lean_KernelException_toMessageData___closed__52(); +lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__52); +l_Lean_KernelException_toMessageData___closed__53 = _init_l_Lean_KernelException_toMessageData___closed__53(); +lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__53); +l_Lean_KernelException_toMessageData___closed__54 = _init_l_Lean_KernelException_toMessageData___closed__54(); +lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__54); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Meta/KAbstract.c b/stage0/stdlib/Lean/Meta/KAbstract.c index b4ad799ff14d..ad1a21455936 100644 --- a/stage0/stdlib/Lean/Meta/KAbstract.c +++ b/stage0/stdlib/Lean/Meta/KAbstract.c @@ -39,6 +39,7 @@ 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_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_399_(uint8_t, uint8_t); static lean_object* l_Lean_Meta_kabstract___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_kabstract___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Meta_Occurrences_contains(lean_object*, lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); uint8_t l___private_Lean_HeadIndex_0__Lean_beqHeadIndex____x40_Lean_HeadIndex___hyg_66_(lean_object*, lean_object*); @@ -6979,7 +6980,6 @@ lean_dec(x_18); x_21 = l_Lean_Meta_kabstract_visit(x_2, x_3, x_14, x_16, x_11, x_15, x_19, x_4, x_5, x_6, x_7, x_20); lean_dec(x_16); lean_dec(x_14); -lean_dec(x_3); if (lean_obj_tag(x_21) == 0) { lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; @@ -7039,7 +7039,6 @@ else { lean_object* x_33; uint8_t x_34; x_33 = lean_box(0); -lean_inc(x_3); x_34 = l___private_Init_MetaTypes_0__Lean_Meta_beqOccurrences____x40_Init_MetaTypes___hyg_1066_(x_3, x_33); if (x_34 == 0) { @@ -7059,7 +7058,6 @@ lean_dec(x_39); x_42 = l_Lean_Meta_kabstract_visit(x_2, x_3, x_35, x_37, x_11, x_36, x_40, x_4, x_5, x_6, x_7, x_41); lean_dec(x_37); lean_dec(x_35); -lean_dec(x_3); if (lean_obj_tag(x_42) == 0) { lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; @@ -7122,7 +7120,6 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); x_54 = l_Lean_Meta_kabstract___closed__1; x_55 = lean_array_push(x_54, x_2); x_56 = lean_expr_abstract(x_11, x_55); @@ -7159,7 +7156,6 @@ lean_dec(x_64); x_67 = l_Lean_Meta_kabstract_visit(x_2, x_3, x_60, x_62, x_57, x_61, x_65, x_4, x_5, x_6, x_7, x_66); lean_dec(x_62); lean_dec(x_60); -lean_dec(x_3); if (lean_obj_tag(x_67) == 0) { lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; @@ -7219,7 +7215,6 @@ else { lean_object* x_78; uint8_t x_79; x_78 = lean_box(0); -lean_inc(x_3); x_79 = l___private_Init_MetaTypes_0__Lean_Meta_beqOccurrences____x40_Init_MetaTypes___hyg_1066_(x_3, x_78); if (x_79 == 0) { @@ -7238,7 +7233,6 @@ lean_dec(x_84); x_87 = l_Lean_Meta_kabstract_visit(x_2, x_3, x_80, x_82, x_57, x_81, x_85, x_4, x_5, x_6, x_7, x_86); lean_dec(x_82); lean_dec(x_80); -lean_dec(x_3); if (lean_obj_tag(x_87) == 0) { lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; @@ -7301,7 +7295,6 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); x_98 = l_Lean_Meta_kabstract___closed__1; x_99 = lean_array_push(x_98, x_2); x_100 = lean_expr_abstract(x_57, x_99); @@ -7316,6 +7309,15 @@ return x_101; } } } +LEAN_EXPORT lean_object* l_Lean_Meta_kabstract___boxed(lean_object* x_1, lean_object* x_2, 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_kabstract(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* initialize_Lean_HeadIndex(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Basic(uint8_t builtin, lean_object*); static bool _G_initialized = false; diff --git a/stage0/stdlib/Lean/Meta/Match/Match.c b/stage0/stdlib/Lean/Meta/Match/Match.c index 24d04b970fab..d36716658781 100644 --- a/stage0/stdlib/Lean/Meta/Match/Match.c +++ b/stage0/stdlib/Lean/Meta/Match/Match.c @@ -85,7 +85,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_pr LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Match_mkMatcherAuxDefinition___spec__2(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_isNextVar___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_withAlts_mkMinorType___spec__1(size_t, size_t, 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_Match___hyg_10661____lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_filterMapM_loop___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processNonVariable___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_solveCnstrs_go___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_throwNonSupported___lambda__1___closed__1; @@ -474,6 +473,7 @@ uint8_t l_List_isEmpty___rarg(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_hasArrayLitPattern___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_traceState___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Match_bootstrap_genMatcherCode; +lean_object* l_instDecidableEqBool___boxed(lean_object*, lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Match_mkMatcher___lambda__17(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Match_State_counterExamples___default; @@ -626,7 +626,6 @@ extern lean_object* l_Lean_Meta_instMonadMetaM; LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_throwNonSupported___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Match_mkMatcher___lambda__18(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_indexOfAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_getUElimPos_x3f___spec__1(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_Match___hyg_10661____lambda__1(uint8_t, uint8_t); LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_expandIntValuePattern(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Match_mkMatcherAuxDefinition___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_Match_withMkMatcherInput___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -28345,35 +28344,11 @@ x_5 = l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6__ return x_5; } } -LEAN_EXPORT uint8_t l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_Match___hyg_10661____lambda__1(uint8_t x_1, uint8_t x_2) { -_start: -{ -if (x_1 == 0) -{ -if (x_2 == 0) -{ -uint8_t x_3; -x_3 = 1; -return x_3; -} -else -{ -uint8_t x_4; -x_4 = 0; -return x_4; -} -} -else -{ -return x_2; -} -} -} static lean_object* _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_Match___hyg_10661____closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_Match___hyg_10661____lambda__1___boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l_instDecidableEqBool___boxed), 2, 0); return x_1; } } @@ -28468,19 +28443,6 @@ x_3 = l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg(x_2, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_Match___hyg_10661____lambda__1___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; -x_3 = lean_unbox(x_1); -lean_dec(x_1); -x_4 = lean_unbox(x_2); -lean_dec(x_2); -x_5 = l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_Match___hyg_10661____lambda__1(x_3, x_4); -x_6 = lean_box(x_5); -return x_6; -} -} LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Match_mkMatcherAuxDefinition___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { diff --git a/stage0/stdlib/Lean/Meta/Tactic/ElimInfo.c b/stage0/stdlib/Lean/Meta/Tactic/ElimInfo.c index 0b15aeda6bba..436896f91046 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/ElimInfo.c +++ b/stage0/stdlib/Lean/Meta/Tactic/ElimInfo.c @@ -86,7 +86,6 @@ static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_29 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_addImplicitTargets_collect___lambda__2___closed__2; -LEAN_EXPORT lean_object* l_Lean_Meta_CustomEliminators_map___default___lambda__1___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_39____closed__3; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___closed__1; static lean_object* l_Lean_Meta_getElimExprInfo___closed__4; @@ -200,7 +199,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_addCustomEliminator(lean_object*, uint8_t, lean_object* l_Lean_ScopedEnvExtension_addScopedEntry___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_getElimExprInfo___lambda__4___closed__2; uint8_t l_Lean_Expr_isSort(lean_object*); -LEAN_EXPORT uint8_t l_Lean_Meta_CustomEliminators_map___default___lambda__1(uint8_t, uint8_t); static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_205____closed__9; lean_object* lean_st_ref_take(lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_3070____closed__1; @@ -256,6 +254,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_getElimExprInfo___lambda__1(lean_object*, l LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_Meta_mkCustomEliminator___spec__10(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_ElimInfo___hyg_2991____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getCustomEliminators___boxed(lean_object*); +lean_object* l_instDecidableEqBool___boxed(lean_object*, 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*); @@ -4893,35 +4892,11 @@ x_2 = l_Lean_mkHashMapImp___rarg(x_1); return x_2; } } -LEAN_EXPORT uint8_t l_Lean_Meta_CustomEliminators_map___default___lambda__1(uint8_t x_1, uint8_t x_2) { -_start: -{ -if (x_1 == 0) -{ -if (x_2 == 0) -{ -uint8_t x_3; -x_3 = 1; -return x_3; -} -else -{ -uint8_t x_4; -x_4 = 0; -return x_4; -} -} -else -{ -return x_2; -} -} -} static lean_object* _init_l_Lean_Meta_CustomEliminators_map___default___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_CustomEliminators_map___default___lambda__1___boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l_instDecidableEqBool___boxed), 2, 0); return x_1; } } @@ -5041,19 +5016,6 @@ lean_dec(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_CustomEliminators_map___default___lambda__1___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; -x_3 = lean_unbox(x_1); -lean_dec(x_1); -x_4 = lean_unbox(x_2); -lean_dec(x_2); -x_5 = l_Lean_Meta_CustomEliminators_map___default___lambda__1(x_3, x_4); -x_6 = lean_box(x_5); -return x_6; -} -} static lean_object* _init_l_Lean_Meta_instInhabitedCustomEliminators___closed__1() { _start: { diff --git a/stage0/stdlib/Lean/Meta/Tactic/LinearArith/Solver.c b/stage0/stdlib/Lean/Meta/Tactic/LinearArith/Solver.c index 5a5031e8824e..28e5e70b112f 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/LinearArith/Solver.c +++ b/stage0/stdlib/Lean/Meta/Tactic/LinearArith/Solver.c @@ -130,6 +130,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instOrdVar; LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_372____spec__4(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3645____closed__1; static lean_object* l_Lean_Meta_Linear_Poly_add_go___closed__1; +lean_object* l_Int_instDecidableEqInt___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Linear_instInhabitedCnstr___closed__1; static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_372____closed__19; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_pickAssignment_x3f(lean_object*, uint8_t, lean_object*, uint8_t); @@ -176,6 +177,7 @@ static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solv LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3645____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_eval_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instReprPoly; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_409____closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_combine_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Linear_Cnstr_isUnsat___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_get___boxed(lean_object*, lean_object*); @@ -250,9 +252,8 @@ static lean_object* l_Lean_Meta_Linear_resolve___closed__1; static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3645____closed__12; lean_object* l_Lean_Rat_floor(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_size(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_409____lambda__1___closed__2; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_409____closed__3; static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_372____spec__2___closed__5; -lean_object* l_Nat_decEq___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_372____closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_add(lean_object*, lean_object*); lean_object* l_Array_back___rarg(lean_object*, lean_object*); @@ -280,7 +281,6 @@ uint8_t lean_int_dec_eq(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3645____closed__9; uint8_t l_Lean_Rat_isInt(lean_object*); extern lean_object* l_Int_instInhabitedInt; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_409____lambda__1___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_82____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_eval_x3f___lambda__1(lean_object*, lean_object*); lean_object* l_Lean_Rat_sub(lean_object*, lean_object*); @@ -292,7 +292,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Linear_CnstrKind_ofNat___boxed(lean_object* lean_object* lean_int_neg(lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3839____boxed(lean_object*, lean_object*); -lean_object* l_Int_decEq___boxed(lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_372____closed__11; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_add_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -300,7 +299,6 @@ static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta LEAN_EXPORT uint8_t l_Lean_Meta_Linear_instDecidableLtVarInstLTVar(lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprAssumptionId____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_2743____boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_409____lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Assignment_val___default; static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3381____closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instReprVar; @@ -1263,45 +1261,39 @@ x_1 = l_Lean_Meta_Linear_instReprPoly___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_409____lambda__1___closed__1() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_409____closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Int_decEq___boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l_Int_instDecidableEqInt___boxed), 2, 0); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_409____lambda__1___closed__2() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_409____closed__2() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Nat_decEq___boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Linear_instDecidableEqVar___boxed), 2, 0); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_409____lambda__1(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_409____lambda__1___closed__1; -x_4 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_409____lambda__1___closed__2; -x_5 = l_instDecidableEqProd___rarg(x_3, x_4, x_1, x_2); -return x_5; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_409____closed__1() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_409____closed__3() { _start: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_409____lambda__1), 2, 0); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_409____closed__1; +x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_409____closed__2; +x_3 = lean_alloc_closure((void*)(l_instDecidableEqProd___rarg), 4, 2); +lean_closure_set(x_3, 0, x_1); +lean_closure_set(x_3, 1, x_2); +return x_3; } } LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_409_(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; uint8_t x_4; -x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_409____closed__1; +x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_409____closed__3; x_4 = l_Array_instDecidableEqArray___rarg(x_3, x_1, x_2); return x_4; } @@ -1320,10 +1312,9 @@ return x_4; LEAN_EXPORT uint8_t l_Lean_Meta_Linear_instDecidableEqPoly(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_3; uint8_t x_4; -x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_409____closed__1; -x_4 = l_Array_instDecidableEqArray___rarg(x_3, x_1, x_2); -return x_4; +uint8_t x_3; +x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_409_(x_1, x_2); +return x_3; } } LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instDecidableEqPoly___boxed(lean_object* x_1, lean_object* x_2) { @@ -3684,30 +3675,29 @@ return x_12; } else { -lean_object* x_13; uint8_t x_14; -x_13 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_409____closed__1; -x_14 = l_Array_instDecidableEqArray___rarg(x_13, x_4, x_8); -if (x_14 == 0) +uint8_t x_13; +x_13 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_409_(x_4, x_8); +if (x_13 == 0) { -uint8_t x_15; -x_15 = 0; -return x_15; +uint8_t x_14; +x_14 = 0; +return x_14; } else { -uint8_t x_16; -x_16 = lean_int_dec_eq(x_5, x_9); -if (x_16 == 0) +uint8_t x_15; +x_15 = lean_int_dec_eq(x_5, x_9); +if (x_15 == 0) { -uint8_t x_17; -x_17 = 0; -return x_17; +uint8_t x_16; +x_16 = 0; +return x_16; } else { -uint8_t x_18; -x_18 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_2832_(x_6, x_10); -return x_18; +uint8_t x_17; +x_17 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_2832_(x_6, x_10); +return x_17; } } } @@ -3764,30 +3754,29 @@ return x_12; } else { -lean_object* x_13; uint8_t x_14; -x_13 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_409____closed__1; -x_14 = l_Array_instDecidableEqArray___rarg(x_13, x_4, x_8); -if (x_14 == 0) +uint8_t x_13; +x_13 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_409_(x_4, x_8); +if (x_13 == 0) { -uint8_t x_15; -x_15 = 0; -return x_15; +uint8_t x_14; +x_14 = 0; +return x_14; } else { -uint8_t x_16; -x_16 = lean_int_dec_eq(x_5, x_9); -if (x_16 == 0) +uint8_t x_15; +x_15 = lean_int_dec_eq(x_5, x_9); +if (x_15 == 0) { -uint8_t x_17; -x_17 = 0; -return x_17; +uint8_t x_16; +x_16 = 0; +return x_16; } else { -uint8_t x_18; -x_18 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3245_(x_6, x_10); -return x_18; +uint8_t x_17; +x_17 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3245_(x_6, x_10); +return x_17; } } } @@ -6146,12 +6135,12 @@ l_Lean_Meta_Linear_instReprPoly___closed__1 = _init_l_Lean_Meta_Linear_instReprP lean_mark_persistent(l_Lean_Meta_Linear_instReprPoly___closed__1); l_Lean_Meta_Linear_instReprPoly = _init_l_Lean_Meta_Linear_instReprPoly(); lean_mark_persistent(l_Lean_Meta_Linear_instReprPoly); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_409____lambda__1___closed__1 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_409____lambda__1___closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_409____lambda__1___closed__1); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_409____lambda__1___closed__2 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_409____lambda__1___closed__2(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_409____lambda__1___closed__2); l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_409____closed__1 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_409____closed__1(); lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_409____closed__1); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_409____closed__2 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_409____closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_409____closed__2); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_409____closed__3 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_409____closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_409____closed__3); l_Lean_Meta_Linear_Poly_getMaxVarCoeff___closed__1 = _init_l_Lean_Meta_Linear_Poly_getMaxVarCoeff___closed__1(); lean_mark_persistent(l_Lean_Meta_Linear_Poly_getMaxVarCoeff___closed__1); l_Lean_Meta_Linear_Poly_add_go___closed__1 = _init_l_Lean_Meta_Linear_Poly_add_go___closed__1(); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Rewrite.c b/stage0/stdlib/Lean/Meta/Tactic/Rewrite.c index 399a41cd62c5..79361383d8c8 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Rewrite.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Rewrite.c @@ -1306,6 +1306,7 @@ lean_inc(x_13); lean_inc(x_3); lean_inc(x_19); x_31 = l_Lean_Meta_kabstract(x_19, x_3, x_21, x_30, x_13, x_14, x_15, x_20); +lean_dec(x_21); if (lean_obj_tag(x_31) == 0) { lean_object* x_32; lean_object* x_33; uint8_t x_34; @@ -1443,6 +1444,7 @@ lean_inc(x_13); lean_inc(x_3); lean_inc(x_19); x_65 = l_Lean_Meta_kabstract(x_19, x_3, x_21, x_64, x_13, x_14, x_15, x_20); +lean_dec(x_21); if (lean_obj_tag(x_65) == 0) { lean_object* x_66; lean_object* x_67; uint8_t x_68; @@ -2244,6 +2246,7 @@ lean_inc(x_13); lean_inc(x_3); lean_inc(x_19); x_31 = l_Lean_Meta_kabstract(x_19, x_3, x_21, x_30, x_13, x_14, x_15, x_20); +lean_dec(x_21); if (lean_obj_tag(x_31) == 0) { lean_object* x_32; lean_object* x_33; uint8_t x_34; @@ -2381,6 +2384,7 @@ lean_inc(x_13); lean_inc(x_3); lean_inc(x_19); x_65 = l_Lean_Meta_kabstract(x_19, x_3, x_21, x_64, x_13, x_14, x_15, x_20); +lean_dec(x_21); if (lean_obj_tag(x_65) == 0) { lean_object* x_66; lean_object* x_67; uint8_t x_68; @@ -3188,6 +3192,7 @@ lean_inc(x_14); lean_inc(x_3); lean_inc(x_20); x_32 = l_Lean_Meta_kabstract(x_20, x_3, x_22, x_31, x_14, x_15, x_16, x_21); +lean_dec(x_22); if (lean_obj_tag(x_32) == 0) { lean_object* x_33; lean_object* x_34; uint8_t x_35; @@ -3327,6 +3332,7 @@ lean_inc(x_14); lean_inc(x_3); lean_inc(x_20); x_66 = l_Lean_Meta_kabstract(x_20, x_3, x_22, x_65, x_14, x_15, x_16, x_21); +lean_dec(x_22); if (lean_obj_tag(x_66) == 0) { lean_object* x_67; lean_object* x_68; uint8_t x_69; @@ -4136,6 +4142,7 @@ lean_inc(x_14); lean_inc(x_3); lean_inc(x_20); x_32 = l_Lean_Meta_kabstract(x_20, x_3, x_22, x_31, x_14, x_15, x_16, x_21); +lean_dec(x_22); if (lean_obj_tag(x_32) == 0) { lean_object* x_33; lean_object* x_34; uint8_t x_35; @@ -4275,6 +4282,7 @@ lean_inc(x_14); lean_inc(x_3); lean_inc(x_20); x_66 = l_Lean_Meta_kabstract(x_20, x_3, x_22, x_65, x_14, x_15, x_16, x_21); +lean_dec(x_22); if (lean_obj_tag(x_66) == 0) { lean_object* x_67; lean_object* x_68; uint8_t x_69; diff --git a/stage0/stdlib/Lean/Parser/Command.c b/stage0/stdlib/Lean/Parser/Command.c index 6f52e9a0bb3d..4b78e9c3d335 100644 --- a/stage0/stdlib/Lean/Parser/Command.c +++ b/stage0/stdlib/Lean/Parser/Command.c @@ -565,6 +565,7 @@ 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; +lean_object* l_instDecidableEqChar___boxed(lean_object*, lean_object*); 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; @@ -2747,7 +2748,6 @@ static lean_object* l_Lean_Parser_Command_declId_formatter___closed__8; lean_object* l_Lean_Parser_Term_leftArrow_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_structExplicitBinder___closed__12; static lean_object* l_Lean_Parser_Command_inductive_parenthesizer___closed__1; -lean_object* l_UInt32_decEq___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_optDeclSig_parenthesizer___closed__10; static lean_object* l___regBuiltin_Lean_Parser_Term_open_declRange___closed__7; static lean_object* l_Lean_Parser_Command_partial_parenthesizer___closed__1; @@ -6903,7 +6903,7 @@ static lean_object* _init_l_Lean_Parser_Command_declId___lambda__1___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_UInt32_decEq___boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l_instDecidableEqChar___boxed), 2, 0); return x_1; } } diff --git a/stage0/stdlib/Lean/Server/References.c b/stage0/stdlib/Lean/Server/References.c index 1fe75f8d1704..18e427de7c58 100644 --- a/stage0/stdlib/Lean/Server/References.c +++ b/stage0/stdlib/Lean/Server/References.c @@ -45,7 +45,6 @@ extern lean_object* l_Lean_declRangeExt; LEAN_EXPORT lean_object* l_Lean_HashMapImp_erase___at_Lean_Server_combineIdents_useConstRepresentatives___spec__6___boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_List_elem___at_Lean_Server_combineIdents_useConstRepresentatives___spec__10(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_References_0__Lean_Server_toJsonIlean____x40_Lean_Server_References___hyg_1464____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l_Lean_Server_dedupReferences___lambda__1(uint8_t, uint8_t); LEAN_EXPORT lean_object* l_Lean_HashMap_insert___at_Lean_Server_combineIdents___spec__2(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_dedupReferences___closed__11; LEAN_EXPORT lean_object* l_Lean_Server_dedupReferences(lean_object*, uint8_t); @@ -116,7 +115,6 @@ LEAN_EXPORT lean_object* l_List_elem___at_Lean_Server_combineIdents_useConstRepr LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Server_References_allRefs___spec__10___boxed(lean_object*, lean_object*); lean_object* l_System_Uri_pathToUri(lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMapImp_moveEntries___at_Lean_Server_combineIdents_useConstRepresentatives___spec__17(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_dedupReferences___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_dedupReferences___spec__11(uint8_t, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_findDeclarationRanges_x3f___at_Lean_Server_RefInfo_toLspRefInfo___spec__2___closed__1; LEAN_EXPORT lean_object* l_Lean_HashMapImp_expand___at_Lean_Server_References_allRefs___spec__3(lean_object*, lean_object*); @@ -235,6 +233,7 @@ lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_ LEAN_EXPORT lean_object* l_Lean_Server_References_removeWorkerRefs(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_erase___at_Lean_Server_References_removeIlean___spec__7(lean_object*, lean_object*); lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonRefInfo___spec__2(size_t, size_t, lean_object*); +lean_object* l_instDecidableEqBool___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Server_dedupReferences___closed__4; static lean_object* l___private_Lean_Server_References_0__Lean_Server_fromJsonIlean____x40_Lean_Server_References___hyg_1325____closed__12; LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Server_combineIdents_useConstRepresentatives___spec__23(lean_object*); @@ -9854,35 +9853,11 @@ return x_4; } } } -LEAN_EXPORT uint8_t l_Lean_Server_dedupReferences___lambda__1(uint8_t x_1, uint8_t x_2) { -_start: -{ -if (x_1 == 0) -{ -if (x_2 == 0) -{ -uint8_t x_3; -x_3 = 1; -return x_3; -} -else -{ -uint8_t x_4; -x_4 = 0; -return x_4; -} -} -else -{ -return x_2; -} -} -} static lean_object* _init_l_Lean_Server_dedupReferences___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_dedupReferences___lambda__1___boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l_instDecidableEqBool___boxed), 2, 0); return x_1; } } @@ -10132,19 +10107,6 @@ lean_dec(x_1); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Server_dedupReferences___lambda__1___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; -x_3 = lean_unbox(x_1); -lean_dec(x_1); -x_4 = lean_unbox(x_2); -lean_dec(x_2); -x_5 = l_Lean_Server_dedupReferences___lambda__1(x_3, x_4); -x_6 = lean_box(x_5); -return x_6; -} -} LEAN_EXPORT lean_object* l_Lean_Server_dedupReferences___boxed(lean_object* x_1, lean_object* x_2) { _start: { diff --git a/stage0/stdlib/Lean/Server/Requests.c b/stage0/stdlib/Lean/Server/Requests.c index 455acd9876f6..a2d67c56a01b 100644 --- a/stage0/stdlib/Lean/Server/Requests.c +++ b/stage0/stdlib/Lean/Server/Requests.c @@ -29,7 +29,6 @@ LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_registerLspRequestHandler LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at_Lean_Server_runHeaderCachingHandlers___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_RequestError_ofException(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_chainLspRequestHandler___spec__1___rarg(lean_object*, lean_object*, lean_object*); -lean_object* l_String_decEq___boxed(lean_object*, lean_object*); uint8_t lean_usize_dec_le(size_t, size_t); lean_object* l_Lean_MessageData_toString(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Server_runHeaderCachingHandlers___spec__5___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -201,6 +200,7 @@ lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_handleLspRequest(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Server_Snapshots_Snapshot_endPos(lean_object*); lean_object* l_Nat_repr(lean_object*); +lean_object* l_instDecidableEqString___boxed(lean_object*, lean_object*); size_t lean_usize_land(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Server_instMonadLiftEIOExceptionRequestM(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_RequestError_invalidParams(lean_object*); @@ -1746,7 +1746,7 @@ static lean_object* _init_l_Lean_Server_initFn____x40_Lean_Server_Requests___hyg _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_String_decEq___boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l_instDecidableEqString___boxed), 2, 0); return x_1; } } diff --git a/stage0/stdlib/Lean/Server/Watchdog.c b/stage0/stdlib/Lean/Server/Watchdog.c index 4eaa3c03dd90..d6d0a27e9a3d 100644 --- a/stage0/stdlib/Lean/Server/Watchdog.c +++ b/stage0/stdlib/Lean/Server/Watchdog.c @@ -79,7 +79,6 @@ static lean_object* l_Lean_Server_Watchdog_tryWriteMessage___closed__52; static lean_object* l_Lean_Server_Watchdog_tryWriteMessage___closed__28; LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Server_Watchdog_handlePrepareCallHierarchy___spec__3___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMapImp_erase___at_Lean_Server_Watchdog_handleCallHierarchyIncomingCalls_collapseSameIncomingCalls___spec__5___boxed(lean_object*, lean_object*); -lean_object* l_String_decEq___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Server_Watchdog_mkLeanServerCapabilities___closed__16; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Watchdog_handleCallHierarchyOutgoingCalls___spec__2(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMapImp_erase___at_Lean_Server_Watchdog_handleCallHierarchyOutgoingCalls_collapseSameOutgoingCalls___spec__5___boxed(lean_object*, lean_object*); @@ -762,6 +761,7 @@ lean_object* l_Lean_SearchPath_findAllWithExt(lean_object*, lean_object*, lean_o static lean_object* l_Lean_Server_Watchdog_handleRename___lambda__1___closed__1; static lean_object* l_Lean_Server_Watchdog_handleRename___lambda__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_initAndRunWatchdogAux___lambda__1(lean_object*, lean_object*, lean_object*); +lean_object* l_instDecidableEqString___boxed(lean_object*, lean_object*); lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonDidChangeWatchedFilesRegistrationOptions____x40_Lean_Data_Lsp_Workspace___hyg_421_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_toStructured_x3f___at_Lean_Server_Watchdog_handleNotification___spec__6(lean_object*); static lean_object* l___private_Lean_Server_Watchdog_0__Lean_Server_Watchdog_fromJsonCallHierarchyItemData____x40_Lean_Server_Watchdog___hyg_3453____closed__13; @@ -19241,7 +19241,7 @@ static lean_object* _init_l_Lean_Server_Watchdog_handleRename___lambda__1___clos _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_String_decEq___boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l_instDecidableEqString___boxed), 2, 0); return x_1; } } From 612d97440b4c0b3da967eaf041687b0f2c468f3a Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Wed, 13 Mar 2024 11:52:57 -0700 Subject: [PATCH 26/32] chore: incorrectly annotated theorems --- src/Init/Core.lean | 9 ++++++--- src/Init/Data/AC.lean | 2 +- src/Init/Data/Nat/Div.lean | 4 ++-- src/Init/WF.lean | 6 +++--- 4 files changed, 12 insertions(+), 9 deletions(-) diff --git a/src/Init/Core.lean b/src/Init/Core.lean index d27d909ac8f4..4b8acf319064 100644 --- a/src/Init/Core.lean +++ b/src/Init/Core.lean @@ -737,13 +737,16 @@ theorem beq_false_of_ne [BEq α] [LawfulBEq α] {a b : α} (h : a ≠ b) : (a == section variable {α β φ : Sort u} {a a' : α} {b b' : β} {c : φ} -theorem HEq.ndrec.{u1, u2} {α : Sort u2} {a : α} {motive : {β : Sort u2} → β → Sort u1} (m : motive a) {β : Sort u2} {b : β} (h : HEq a b) : motive b := +/-- Non-dependent recursor for `HEq` -/ +noncomputable def HEq.ndrec.{u1, u2} {α : Sort u2} {a : α} {motive : {β : Sort u2} → β → Sort u1} (m : motive a) {β : Sort u2} {b : β} (h : HEq a b) : motive b := h.rec m -theorem HEq.ndrecOn.{u1, u2} {α : Sort u2} {a : α} {motive : {β : Sort u2} → β → Sort u1} {β : Sort u2} {b : β} (h : HEq a b) (m : motive a) : motive b := +/-- `HEq.ndrec` variant -/ +noncomputable def HEq.ndrecOn.{u1, u2} {α : Sort u2} {a : α} {motive : {β : Sort u2} → β → Sort u1} {β : Sort u2} {b : β} (h : HEq a b) (m : motive a) : motive b := h.rec m -theorem HEq.elim {α : Sort u} {a : α} {p : α → Sort v} {b : α} (h₁ : HEq a b) (h₂ : p a) : p b := +/-- `HEq.ndrec` variant -/ +noncomputable def HEq.elim {α : Sort u} {a : α} {p : α → Sort v} {b : α} (h₁ : HEq a b) (h₂ : p a) : p b := eq_of_heq h₁ ▸ h₂ theorem HEq.subst {p : (T : Sort u) → T → Prop} (h₁ : HEq a b) (h₂ : p α a) : p β b := diff --git a/src/Init/Data/AC.lean b/src/Init/Data/AC.lean index a05d15a9a94f..ffae0489a5bb 100644 --- a/src/Init/Data/AC.lean +++ b/src/Init/Data/AC.lean @@ -106,7 +106,7 @@ def norm [info : ContextInformation α] (ctx : α) (e : Expr) : List Nat := let xs := if info.isComm ctx then sort xs else xs if info.isIdem ctx then mergeIdem xs else xs -theorem List.two_step_induction +noncomputable def List.two_step_induction {motive : List Nat → Sort u} (l : List Nat) (empty : motive []) diff --git a/src/Init/Data/Nat/Div.lean b/src/Init/Data/Nat/Div.lean index a8fed3caf2ae..9407c4284507 100644 --- a/src/Init/Data/Nat/Div.lean +++ b/src/Init/Data/Nat/Div.lean @@ -35,7 +35,7 @@ theorem div_eq (x y : Nat) : x / y = if 0 < y ∧ y ≤ x then (x - y) / y + 1 e rw [Nat.div] rfl -theorem div.inductionOn.{u} +def div.inductionOn.{u} {motive : Nat → Nat → Sort u} (x y : Nat) (ind : ∀ x y, 0 < y ∧ y ≤ x → motive (x - y) y → motive x y) @@ -102,7 +102,7 @@ protected theorem modCore_eq_mod (x y : Nat) : Nat.modCore x y = x % y := by theorem mod_eq (x y : Nat) : x % y = if 0 < y ∧ y ≤ x then (x - y) % y else x := by rw [←Nat.modCore_eq_mod, ←Nat.modCore_eq_mod, Nat.modCore] -theorem mod.inductionOn.{u} +def mod.inductionOn.{u} {motive : Nat → Nat → Sort u} (x y : Nat) (ind : ∀ x y, 0 < y ∧ y ≤ x → motive (x - y) y → motive x y) diff --git a/src/Init/WF.lean b/src/Init/WF.lean index 9076f5542740..da0f9c050a2c 100644 --- a/src/Init/WF.lean +++ b/src/Init/WF.lean @@ -45,7 +45,7 @@ def apply {α : Sort u} {r : α → α → Prop} (wf : WellFounded r) (a : α) : section variable {α : Sort u} {r : α → α → Prop} (hwf : WellFounded r) -theorem recursion {C : α → Sort v} (a : α) (h : ∀ x, (∀ y, r y x → C y) → C x) : C a := by +noncomputable def recursion {C : α → Sort v} (a : α) (h : ∀ x, (∀ y, r y x → C y) → C x) : C a := by induction (apply hwf a) with | intro x₁ _ ih => exact h x₁ ih @@ -166,13 +166,13 @@ def lt_wfRel : WellFoundedRelation Nat where | Or.inl e => subst e; assumption | Or.inr e => exact Acc.inv ih e -protected theorem strongInductionOn +protected noncomputable def strongInductionOn {motive : Nat → Sort u} (n : Nat) (ind : ∀ n, (∀ m, m < n → motive m) → motive n) : motive n := Nat.lt_wfRel.wf.fix ind n -protected theorem caseStrongInductionOn +protected noncomputable def caseStrongInductionOn {motive : Nat → Sort u} (a : Nat) (zero : motive 0) From 8d2adf521d2b7636347a5b01bfe473bf0fcfaf31 Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Tue, 12 Mar 2024 17:53:04 -0700 Subject: [PATCH 27/32] feat: allow duplicate theorems to be imported --- RELEASES.md | 2 ++ src/Lean/Data/HashMap.lean | 34 +++++++++++++++++++++++++++++---- src/Lean/Environment.lean | 37 ++++++++++++++++++++++++++++++++---- tests/pkg/misc/Misc/Boo.lean | 2 ++ tests/pkg/misc/Misc/Foo.lean | 2 ++ 5 files changed, 69 insertions(+), 8 deletions(-) diff --git a/RELEASES.md b/RELEASES.md index f23616684e54..ee28c1a13f5b 100644 --- a/RELEASES.md +++ b/RELEASES.md @@ -13,6 +13,8 @@ v4.8.0 (development in progress) * Lean now generates an error if the type of a theorem is **not** a proposition. +* Importing two different files containing proofs of the same theorem is no longer considered an error. This feature is particularly useful for theorems that are automatically generated on demand (e.g., equational theorems). + * New command `derive_functinal_induction`: Derived from the definition of a (possibly mutually) recursive function diff --git a/src/Lean/Data/HashMap.lean b/src/Lean/Data/HashMap.lean index 7e7dc52344f8..053593076b09 100644 --- a/src/Lean/Data/HashMap.lean +++ b/src/Lean/Data/HashMap.lean @@ -116,6 +116,22 @@ def expand [Hashable α] (size : Nat) (buckets : HashMapBucket α β) : HashMapI else (expand size' buckets', false) +@[inline] def insertIfNew [beq : BEq α] [Hashable α] (m : HashMapImp α β) (a : α) (b : β) : HashMapImp α β × Option β := + match m with + | ⟨size, buckets⟩ => + let ⟨i, h⟩ := mkIdx (hash a) buckets.property + let bkt := buckets.val[i] + if let some b := bkt.find? a then + (m, some b) + else + let size' := size + 1 + let buckets' := buckets.update i (AssocList.cons a b bkt) h + if numBucketsForCapacity size' ≤ buckets.val.size then + ({ size := size', buckets := buckets' }, none) + else + (expand size' buckets', none) + + def erase [BEq α] [Hashable α] (m : HashMapImp α β) (a : α) : HashMapImp α β := match m with | ⟨ size, buckets ⟩ => @@ -125,9 +141,10 @@ def erase [BEq α] [Hashable α] (m : HashMapImp α β) (a : α) : HashMapImp α else m inductive WellFormed [BEq α] [Hashable α] : HashMapImp α β → Prop where - | mkWff : ∀ n, WellFormed (mkHashMapImp n) - | insertWff : ∀ m a b, WellFormed m → WellFormed (insert m a b |>.1) - | eraseWff : ∀ m a, WellFormed m → WellFormed (erase m a) + | mkWff : ∀ n, WellFormed (mkHashMapImp n) + | insertWff : ∀ m a b, WellFormed m → WellFormed (insert m a b |>.1) + | insertIfNewWff : ∀ m a b, WellFormed m → WellFormed (insertIfNew m a b |>.1) + | eraseWff : ∀ m a, WellFormed m → WellFormed (erase m a) end HashMapImp @@ -156,13 +173,22 @@ def insert (m : HashMap α β) (a : α) (b : β) : HashMap α β := match h:m.insert a b with | (m', _) => ⟨ m', by have aux := WellFormed.insertWff m a b hw; rw [h] at aux; assumption ⟩ -/-- Similar to `insert`, but also returns a Boolean flad indicating whether an existing entry has been replaced with `a -> b`. -/ +/-- Similar to `insert`, but also returns a Boolean flag indicating whether an existing entry has been replaced with `a -> b`. -/ def insert' (m : HashMap α β) (a : α) (b : β) : HashMap α β × Bool := match m with | ⟨ m, hw ⟩ => match h:m.insert a b with | (m', replaced) => (⟨ m', by have aux := WellFormed.insertWff m a b hw; rw [h] at aux; assumption ⟩, replaced) +/-- +Similar to `insert`, but returns `some old` if the map already had an entry `α → old`. +If the result is `some old`, the the resulting map is equal to `m`. -/ +def insertIfNew (m : HashMap α β) (a : α) (b : β) : HashMap α β × Option β := + match m with + | ⟨ m, hw ⟩ => + match h:m.insertIfNew a b with + | (m', old) => (⟨ m', by have aux := WellFormed.insertIfNewWff m a b hw; rw [h] at aux; assumption ⟩, old) + @[inline] def erase (m : HashMap α β) (a : α) : HashMap α β := match m with | ⟨ m, hw ⟩ => ⟨ m.erase a, WellFormed.eraseWff m a hw ⟩ diff --git a/src/Lean/Environment.lean b/src/Lean/Environment.lean index 8c5f6dd6ad23..7d4062ad6288 100644 --- a/src/Lean/Environment.lean +++ b/src/Lean/Environment.lean @@ -770,6 +770,33 @@ partial def importModulesCore (imports : Array Import) : ImportStateM Unit := do moduleNames := s.moduleNames.push i.module } +/-- +Return `true` if `cinfo₁` and `cinfo₂` are theorems with the same name, universe parameters, +and types. We allow different modules to prove the same theorem. + +Motivation: We want to generate equational theorems on demand and potentially +in different files, and we want them to have non-private names. +We may add support for other kinds of definitions in the future. For now, theorems are +sufficient for our purposes. + +We may have to revise this design decision and eagerly generate equational theorems when +we implement the module system. + +Remark: we do not check whether the theorem `value` field match. This feature is useful and +ensures the proofs for equational theorems do not need to be identical. This decision +relies on the fact that theorem types are propositions, we have proof irrelevance, +and theorems are (mostly) opaque in Lean. For `Acc.rec`, we may unfold theorems +during type-checking, but we are assuming this is not an issue in practice, +and we are planning to address this issue in the future. +-/ +private def equivInfo (cinfo₁ cinfo₂ : ConstantInfo) : Bool := Id.run do + let .thmInfo tval₁ := cinfo₁ | false + let .thmInfo tval₂ := cinfo₂ | false + return tval₁.name == tval₂.name + && tval₁.type == tval₂.type + && tval₁.levelParams == tval₂.levelParams + && tval₁.all == tval₂.all + /-- Construct environment from `importModulesCore` results. @@ -788,11 +815,13 @@ def finalizeImport (s : ImportState) (imports : Array Import) (opts : Options) ( for h:modIdx in [0:s.moduleData.size] do let mod := s.moduleData[modIdx]'h.upper for cname in mod.constNames, cinfo in mod.constants do - match constantMap.insert' cname cinfo with - | (constantMap', replaced) => + match constantMap.insertIfNew cname cinfo with + | (constantMap', cinfoPrev?) => constantMap := constantMap' - if replaced then - throwAlreadyImported s const2ModIdx modIdx cname + if let some cinfoPrev := cinfoPrev? then + -- Recall that the map has not been modified when `cinfoPrev? = some _`. + unless equivInfo cinfoPrev cinfo do + throwAlreadyImported s const2ModIdx modIdx cname const2ModIdx := const2ModIdx.insert cname modIdx for cname in mod.extraConstNames do const2ModIdx := const2ModIdx.insert cname modIdx diff --git a/tests/pkg/misc/Misc/Boo.lean b/tests/pkg/misc/Misc/Boo.lean index 406fb8e0c2b8..7fc0c25fa391 100644 --- a/tests/pkg/misc/Misc/Boo.lean +++ b/tests/pkg/misc/Misc/Boo.lean @@ -1,3 +1,5 @@ local infix:50 " ≺ " => LE.le #check 1 ≺ 2 + +theorem simple : 10 = 10 := rfl diff --git a/tests/pkg/misc/Misc/Foo.lean b/tests/pkg/misc/Misc/Foo.lean index 918d229608ce..ffa0d91271c2 100644 --- a/tests/pkg/misc/Misc/Foo.lean +++ b/tests/pkg/misc/Misc/Foo.lean @@ -16,3 +16,5 @@ local macro "my_refl" : tactic => `(tactic| rfl) def f (x y : Nat) (_h : x = y := by my_refl) := x + +theorem simple : 10 = 10 := by decide From 2c8fd7fb95331aa5e1c1e1155cfc23e84eafff3a Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Wed, 13 Mar 2024 20:09:15 -0700 Subject: [PATCH 28/32] chore: avoid reserved name TODO: update state0 and cleanup --- src/Lean/Elab/Declaration.lean | 3 ++- src/Lean/Elab/DefView.lean | 7 +++++-- src/Lean/Linter/MissingDocs.lean | 2 +- src/Lean/Parser/Command.lean | 6 ++++-- stage0/src/stdlib_flags.h | 2 +- 5 files changed, 13 insertions(+), 7 deletions(-) diff --git a/src/Lean/Elab/Declaration.lean b/src/Lean/Elab/Declaration.lean index c8af67d77a30..90e2cf682b12 100644 --- a/src/Lean/Elab/Declaration.lean +++ b/src/Lean/Elab/Declaration.lean @@ -42,7 +42,8 @@ private def isNamedDef (stx : Syntax) : Bool := let decl := stx[1] let k := decl.getKind k == ``Lean.Parser.Command.abbrev || - k == ``Lean.Parser.Command.def || + k == ``Lean.Parser.Command.definition || + k == ``Lean.Parser.Command.def || -- TODO: delete k == ``Lean.Parser.Command.theorem || k == ``Lean.Parser.Command.opaque || k == ``Lean.Parser.Command.axiom || diff --git a/src/Lean/Elab/DefView.lean b/src/Lean/Elab/DefView.lean index 246b9987f775..71c1f9242986 100644 --- a/src/Lean/Elab/DefView.lean +++ b/src/Lean/Elab/DefView.lean @@ -142,7 +142,8 @@ def mkDefViewOfExample (modifiers : Modifiers) (stx : Syntax) : DefView := def isDefLike (stx : Syntax) : Bool := let declKind := stx.getKind declKind == ``Parser.Command.abbrev || - declKind == ``Parser.Command.def || + declKind == ``Parser.Command.definition || + declKind == ``Parser.Command.def || -- TODO: delete declKind == ``Parser.Command.theorem || declKind == ``Parser.Command.opaque || declKind == ``Parser.Command.instance || @@ -152,7 +153,9 @@ def mkDefView (modifiers : Modifiers) (stx : Syntax) : CommandElabM DefView := let declKind := stx.getKind if declKind == ``Parser.Command.«abbrev» then return mkDefViewOfAbbrev modifiers stx - else if declKind == ``Parser.Command.def then + else if declKind == ``Parser.Command.definition then + return mkDefViewOfDef modifiers stx + else if declKind == ``Parser.Command.def then -- TODO: delete return mkDefViewOfDef modifiers stx else if declKind == ``Parser.Command.theorem then return mkDefViewOfTheorem modifiers stx diff --git a/src/Lean/Linter/MissingDocs.lean b/src/Lean/Linter/MissingDocs.lean index e9c6e97882be..81930a1a7c12 100644 --- a/src/Lean/Linter/MissingDocs.lean +++ b/src/Lean/Linter/MissingDocs.lean @@ -123,7 +123,7 @@ def declModifiersPubNoDoc (mods : Syntax) : Bool := def lintDeclHead (k : SyntaxNodeKind) (id : Syntax) : CommandElabM Unit := do if k == ``«abbrev» then lintNamed id "public abbrev" - else if k == ``«def» then lintNamed id "public def" + else if k == ``definition then lintNamed id "public def" else if k == ``«opaque» then lintNamed id "public opaque" else if k == ``«axiom» then lintNamed id "public axiom" else if k == ``«inductive» then lintNamed id "public inductive" diff --git a/src/Lean/Parser/Command.lean b/src/Lean/Parser/Command.lean index 1558d097d8c9..995d55d43fb4 100644 --- a/src/Lean/Parser/Command.lean +++ b/src/Lean/Parser/Command.lean @@ -114,7 +114,9 @@ def «abbrev» := leading_parser "abbrev " >> declId >> ppIndent optDeclSig >> declVal def optDefDeriving := optional (ppDedent ppLine >> atomic ("deriving " >> notSymbol "instance") >> sepBy1 ident ", ") -def «def» := leading_parser +def definition := leading_parser + "def " >> recover declId skipUntilWsOrDelim >> ppIndent optDeclSig >> declVal >> optDefDeriving +def «def» := leading_parser -- TODO: delete "def " >> recover declId skipUntilWsOrDelim >> ppIndent optDeclSig >> declVal >> optDefDeriving def «theorem» := leading_parser "theorem " >> recover declId skipUntilWsOrDelim >> ppIndent declSig >> declVal @@ -198,7 +200,7 @@ def «structure» := leading_parser optDeriving @[builtin_command_parser] def declaration := leading_parser declModifiers false >> - («abbrev» <|> «def» <|> «theorem» <|> «opaque» <|> «instance» <|> «axiom» <|> «example» <|> + («abbrev» <|> definition <|> «def» <|> «theorem» <|> «opaque» <|> «instance» <|> «axiom» <|> «example» <|> «inductive» <|> classInductive <|> «structure») @[builtin_command_parser] def «deriving» := leading_parser "deriving " >> "instance " >> derivingClasses >> " for " >> sepBy1 (recover ident skip) ", " diff --git a/stage0/src/stdlib_flags.h b/stage0/src/stdlib_flags.h index 0699845ba452..658ab0874e68 100644 --- a/stage0/src/stdlib_flags.h +++ b/stage0/src/stdlib_flags.h @@ -8,7 +8,7 @@ options get_default_options() { // switch to `true` for ABI-breaking changes affecting meta code opts = opts.update({"interpreter", "prefer_native"}, false); // switch to `true` for changing built-in parsers used in quotations - opts = opts.update({"internal", "parseQuotWithCurrentStage"}, false); + opts = opts.update({"internal", "parseQuotWithCurrentStage"}, true); // toggling `parseQuotWithCurrentStage` may also require toggling the following option if macros/syntax // with custom precheck hooks were affected opts = opts.update({"quotPrecheck"}, true); From 653eb5f66ef4026fa623ffeeaff349e36592fb0c Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Wed, 13 Mar 2024 20:44:10 -0700 Subject: [PATCH 29/32] chore: update stage0 --- stage0/src/stdlib_flags.h | 2 +- stage0/stdlib/Init/Data/Nat/Div.c | 68 + stage0/stdlib/Init/NotationExtra.c | 435 +- stage0/stdlib/Lean/Compiler/CSimpAttr.c | 6 +- stage0/stdlib/Lean/Compiler/LCNF/FloatLetIn.c | 2 +- stage0/stdlib/Lean/Compiler/LCNF/JoinPoints.c | 2 +- stage0/stdlib/Lean/Data/HashMap.c | 308 +- stage0/stdlib/Lean/Elab/AuxDef.c | 63 +- stage0/stdlib/Lean/Elab/Declaration.c | 301 +- stage0/stdlib/Lean/Elab/DefView.c | 350 +- stage0/stdlib/Lean/Elab/Deriving/BEq.c | 217 +- stage0/stdlib/Lean/Elab/Deriving/DecEq.c | 91 +- stage0/stdlib/Lean/Elab/Deriving/FromToJson.c | 229 +- stage0/stdlib/Lean/Elab/Deriving/Hashable.c | 179 +- stage0/stdlib/Lean/Elab/Deriving/Ord.c | 159 +- stage0/stdlib/Lean/Elab/Deriving/Repr.c | 165 +- stage0/stdlib/Lean/Elab/Deriving/TypeName.c | 389 +- stage0/stdlib/Lean/Elab/Inductive.c | 137 +- stage0/stdlib/Lean/Elab/PreDefinition/Eqns.c | 2 +- .../Lean/Elab/PreDefinition/Structural/Eqns.c | 2 +- .../stdlib/Lean/Elab/PreDefinition/WF/Eqns.c | 2 +- stage0/stdlib/Lean/Elab/Syntax.c | 305 +- stage0/stdlib/Lean/Elab/Tactic/Omega/Core.c | 2847 +- .../stdlib/Lean/Elab/Tactic/Omega/Frontend.c | 25925 ++++++++-------- stage0/stdlib/Lean/Elab/Tactic/Omega/OmegaM.c | 5112 +-- stage0/stdlib/Lean/Environment.c | 1152 +- stage0/stdlib/Lean/Linter/Builtin.c | 2 +- stage0/stdlib/Lean/Linter/MissingDocs.c | 2 +- stage0/stdlib/Lean/Meta.c | 6 +- stage0/stdlib/Lean/Meta/Canonicalizer.c | 5364 ++++ stage0/stdlib/Lean/Meta/Tactic/AC/Main.c | 2 +- stage0/stdlib/Lean/Parser/Command.c | 1471 +- .../Delaborator/TopDownAnalyze.c | 2 +- stage0/stdlib/Lean/Replay.c | 2 +- .../Lean/Server/FileWorker/RequestHandling.c | 1868 +- stage0/stdlib/Lean/Server/References.c | 2 +- 36 files changed, 27543 insertions(+), 19628 deletions(-) create mode 100644 stage0/stdlib/Lean/Meta/Canonicalizer.c diff --git a/stage0/src/stdlib_flags.h b/stage0/src/stdlib_flags.h index 658ab0874e68..0699845ba452 100644 --- a/stage0/src/stdlib_flags.h +++ b/stage0/src/stdlib_flags.h @@ -8,7 +8,7 @@ options get_default_options() { // switch to `true` for ABI-breaking changes affecting meta code opts = opts.update({"interpreter", "prefer_native"}, false); // switch to `true` for changing built-in parsers used in quotations - opts = opts.update({"internal", "parseQuotWithCurrentStage"}, true); + opts = opts.update({"internal", "parseQuotWithCurrentStage"}, false); // toggling `parseQuotWithCurrentStage` may also require toggling the following option if macros/syntax // with custom precheck hooks were affected opts = opts.update({"quotPrecheck"}, true); diff --git a/stage0/stdlib/Init/Data/Nat/Div.c b/stage0/stdlib/Init/Data/Nat/Div.c index 83578432be96..dea23efa5bbd 100644 --- a/stage0/stdlib/Init/Data/Nat/Div.c +++ b/stage0/stdlib/Init/Data/Nat/Div.c @@ -15,14 +15,21 @@ extern "C" { #endif LEAN_EXPORT lean_object* l_Nat_instDivNat; LEAN_EXPORT lean_object* l_Nat_modCore___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Nat_mod_inductionOn(lean_object*); +LEAN_EXPORT lean_object* l_Nat_div_inductionOn(lean_object*); LEAN_EXPORT lean_object* l_Nat_instModNat; LEAN_EXPORT lean_object* l_Nat_div___boxed(lean_object*, lean_object*); lean_object* lean_nat_div(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_instDvdNat; +LEAN_EXPORT lean_object* l_Nat_div_inductionOn___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Nat_mod_inductionOn___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_mod(lean_object*, lean_object*); +uint8_t lean_nat_dec_lt(lean_object*, lean_object*); lean_object* lean_nat_mod(lean_object*, lean_object*); static lean_object* l_Nat_instModNat___closed__1; +lean_object* lean_nat_sub(lean_object*, lean_object*); static lean_object* l_Nat_instDivNat___closed__1; +uint8_t lean_nat_dec_le(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_mod___boxed(lean_object*, lean_object*); static lean_object* _init_l_Nat_instDvdNat() { _start: @@ -58,6 +65,51 @@ x_1 = l_Nat_instDivNat___closed__1; return x_1; } } +LEAN_EXPORT lean_object* l_Nat_div_inductionOn___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; uint8_t x_6; +x_5 = lean_unsigned_to_nat(0u); +x_6 = lean_nat_dec_lt(x_5, x_2); +if (x_6 == 0) +{ +lean_object* x_7; +lean_dec(x_3); +x_7 = lean_apply_3(x_4, x_1, x_2, lean_box(0)); +return x_7; +} +else +{ +uint8_t x_8; +x_8 = lean_nat_dec_le(x_2, x_1); +if (x_8 == 0) +{ +lean_object* x_9; +lean_dec(x_3); +x_9 = lean_apply_3(x_4, x_1, x_2, lean_box(0)); +return x_9; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_nat_sub(x_1, x_2); +lean_inc(x_3); +lean_inc(x_2); +x_11 = l_Nat_div_inductionOn___rarg(x_10, x_2, x_3, x_4); +x_12 = lean_apply_4(x_3, x_1, x_2, lean_box(0), x_11); +return x_12; +} +} +} +} +LEAN_EXPORT lean_object* l_Nat_div_inductionOn(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Nat_div_inductionOn___rarg), 4, 0); +return x_2; +} +} LEAN_EXPORT lean_object* l_Nat_modCore___boxed(lean_object* x_1, lean_object* x_2) { _start: { @@ -94,6 +146,22 @@ x_1 = l_Nat_instModNat___closed__1; return x_1; } } +LEAN_EXPORT lean_object* l_Nat_mod_inductionOn___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_Nat_div_inductionOn___rarg(x_1, x_2, x_3, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Nat_mod_inductionOn(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Nat_mod_inductionOn___rarg), 4, 0); +return x_2; +} +} lean_object* initialize_Init_WF(uint8_t builtin, lean_object*); lean_object* initialize_Init_WFTactics(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_Nat_Basic(uint8_t builtin, lean_object*); diff --git a/stage0/stdlib/Init/NotationExtra.c b/stage0/stdlib/Init/NotationExtra.c index b738ed871cd4..5e28ceaab107 100644 --- a/stage0/stdlib/Init/NotationExtra.c +++ b/stage0/stdlib/Init/NotationExtra.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Init.NotationExtra -// Imports: Init.Meta Init.Data.Array.Subarray Init.Data.ToString +// Imports: Init.Meta Init.Data.Array.Subarray Init.Data.ToString Init.Conv #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -134,6 +134,7 @@ static lean_object* l_term___xd7____1___closed__5; static lean_object* l_tacticFunext_________closed__13; static lean_object* l___aux__Init__NotationExtra______macroRules__Lean__Parser__Term__letI__1___closed__1; static lean_object* l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__solve__1___spec__3___closed__1; +static lean_object* l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__8; 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; @@ -152,6 +153,7 @@ static lean_object* l___aux__Init__NotationExtra______macroRules__Lean__Parser__ 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__22; lean_object* lean_array_fget(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_convCalc__; 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; LEAN_EXPORT lean_object* l_unexpandMkArray7(lean_object*, lean_object*, lean_object*); @@ -193,6 +195,7 @@ static lean_object* l_Lean_unbracketedExplicitBinders___closed__12; static lean_object* l_Lean_unifConstraintElem___closed__2; static lean_object* l_Lean_unifConstraint___closed__15; LEAN_EXPORT lean_object* l_Lean_expandBrackedBindersAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_convCalc_____closed__1; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__4; static lean_object* l_solve___closed__9; static lean_object* l_unexpandListToArray___closed__1; @@ -210,6 +213,7 @@ lean_object* l_Lean_Syntax_node5(lean_object*, lean_object*, lean_object*, lean_ static lean_object* l_cdotTk___closed__6; uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_unexpandGetElem_x21(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___aux__Init__NotationExtra______macroRules__Lean__doElemWhile___x3a__Do____1___closed__7; static lean_object* l_Lean_unifConstraintElem___closed__6; static lean_object* l_Lean_Parser_Command_classAbbrev___closed__27; @@ -247,6 +251,7 @@ static lean_object* l_term___xd7____1___closed__4; static lean_object* l_term_u03a3___x2c_____closed__8; 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_u03a3_x27___x2c_____closed__6; +static lean_object* l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__7; static lean_object* l_Lean_explicitBinders___closed__5; static lean_object* l_calcSteps___closed__9; static lean_object* l___aux__Init__NotationExtra______macroRules__Lean__Parser__Command__classAbbrev__1___lambda__3___closed__5; @@ -388,6 +393,8 @@ static lean_object* l_Lean_expandExplicitBindersAux_loop___closed__1; LEAN_EXPORT lean_object* l_term_x7b___x7d; static lean_object* l___aux__Init__NotationExtra______macroRules__solve__1___closed__4; static lean_object* l___aux__Init__NotationExtra______macroRules__term_x7b___x7d__1___lambda__1___closed__5; +static lean_object* l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__1; +static lean_object* l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__3; 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*); @@ -434,7 +441,6 @@ 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_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_doElemWhile___x3a__Do_____closed__8; static lean_object* l_calcSteps___closed__14; static lean_object* l_Lean_termMacro_x2etrace_x5b___x5d_____closed__4; @@ -467,6 +473,7 @@ static lean_object* l___aux__Init__NotationExtra______macroRules__term_x7b___x7d 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___aux__Init__NotationExtra______macroRules__convCalc____1___closed__5; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemWhile___x3a__Do____1___closed__4; static lean_object* l_unexpandEqNDRec___closed__1; LEAN_EXPORT lean_object* l_unexpandMkArray0(lean_object*); @@ -481,9 +488,11 @@ 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*); 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_convCalc_____closed__3; 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; +static lean_object* l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__6; LEAN_EXPORT lean_object* l_unexpandUnit(lean_object*); static lean_object* l_Lean_command____Unif__hint________Where___x7c___x2d_u22a2_____closed__19; lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); @@ -529,7 +538,7 @@ lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); static lean_object* l_Lean_unbracketedExplicitBinders___closed__14; static lean_object* l___aux__Init__NotationExtra______macroRules__Lean__Parser__Term__haveI__1___lambda__1___closed__2; 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_convCalc_____closed__2; static lean_object* l_unexpandListNil___rarg___closed__2; uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); static lean_object* l_Lean_unifConstraint___closed__9; @@ -617,12 +626,14 @@ static lean_object* l_Lean_bracketedExplicitBinders___closed__5; static lean_object* l_Lean_command____Unif__hint________Where___x7c___x2d_u22a2_____closed__50; 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___aux__Init__NotationExtra______macroRules__convCalc____1___closed__4; static lean_object* l_calcSteps___closed__13; static lean_object* l_Lean_unifConstraintElem___closed__3; static lean_object* l_Lean_command____Unif__hint________Where___x7c___x2d_u22a2_____closed__13; 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; +static lean_object* l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__9; LEAN_EXPORT lean_object* l___aux__Init__NotationExtra______macroRules__Lean__Parser__Term__letI__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_Lean_Parser_Command_classAbbrev___closed__13; static lean_object* l___aux__Init__NotationExtra______macroRules__Lean__Parser__Command__classAbbrev__1___lambda__3___closed__4; @@ -657,6 +668,7 @@ LEAN_EXPORT lean_object* l___aux__Init__NotationExtra______macroRules__term___xd static lean_object* l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__5; static lean_object* l___aux__Init__NotationExtra______macroRules__term_x7b___x7d__1___closed__6; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__18; +static lean_object* l___aux__Init__NotationExtra______macroRules__convCalc____1___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__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_unexpandGetElem_x21___closed__1; static lean_object* l_term___xd7_x27_____closed__1; @@ -720,6 +732,7 @@ lean_object* l_Lean_Syntax_node1(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkSepArray(lean_object*, lean_object*); 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___aux__Init__NotationExtra______macroRules__convCalc____1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instForInLoopUnit___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_binderIdent; static lean_object* l_Lean_doElemWhile__Do_____closed__3; @@ -794,7 +807,6 @@ 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; LEAN_EXPORT lean_object* l_Lean_doElemWhile___x3a__Do__; 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; @@ -852,7 +864,6 @@ static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__ter lean_object* lean_string_append(lean_object*, lean_object*); static lean_object* l_term_x7b___x7d___closed__5; 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_command____Unif__hint________Where___x7c___x2d_u22a2_____closed__18; 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_Lean_doElemRepeat_____closed__4; @@ -922,7 +933,6 @@ static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__ter LEAN_EXPORT lean_object* l_unexpandMkArray8(lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__NotationExtra______macroRules__term_u03a3_x27___x2c____1___closed__1; static lean_object* l_solve___closed__4; -static lean_object* l___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__7; 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_doElemWhile__Do_____closed__4; static lean_object* l_Lean_unifConstraint___closed__8; @@ -6898,6 +6908,199 @@ x_1 = l_calcTactic___closed__5; return x_1; } } +static lean_object* _init_l_convCalc_____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("convCalc_", 9); +return x_1; +} +} +static lean_object* _init_l_convCalc_____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_convCalc_____closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_convCalc_____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_convCalc_____closed__2; +x_2 = lean_unsigned_to_nat(1022u); +x_3 = l_calcTactic___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_convCalc__() { +_start: +{ +lean_object* x_1; +x_1 = l_convCalc_____closed__3; +return x_1; +} +} +static lean_object* _init_l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Tactic", 6); +return x_1; +} +} +static lean_object* _init_l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Conv", 4); +return x_1; +} +} +static lean_object* _init_l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("nestedTactic", 12); +return x_1; +} +} +static lean_object* _init_l___aux__Init__NotationExtra______macroRules__convCalc____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_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___aux__Init__NotationExtra______macroRules__convCalc____1___closed__1; +x_4 = l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__2; +x_5 = l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__3; +x_6 = l_Lean_Name_mkStr5(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +static lean_object* _init_l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("tactic", 6); +return x_1; +} +} +static lean_object* _init_l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("tacticSeq", 9); +return x_1; +} +} +static lean_object* _init_l___aux__Init__NotationExtra______macroRules__convCalc____1___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_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___aux__Init__NotationExtra______macroRules__convCalc____1___closed__1; +x_4 = l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__6; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("tacticSeq1Indented", 18); +return x_1; +} +} +static lean_object* _init_l___aux__Init__NotationExtra______macroRules__convCalc____1___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_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___aux__Init__NotationExtra______macroRules__convCalc____1___closed__1; +x_4 = l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__8; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l___aux__Init__NotationExtra______macroRules__convCalc____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_convCalc_____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; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; 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_8 = lean_unsigned_to_nat(0u); +x_9 = l_Lean_Syntax_getArg(x_1, x_8); +x_10 = lean_unsigned_to_nat(1u); +x_11 = l_Lean_Syntax_getArg(x_1, x_10); +lean_dec(x_1); +x_12 = lean_ctor_get(x_2, 5); +lean_inc(x_12); +lean_dec(x_2); +x_13 = 0; +x_14 = l_Lean_SourceInfo_fromRef(x_12, x_13); +x_15 = l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__5; +lean_inc(x_14); +x_16 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +x_17 = l_Lean_expandExplicitBindersAux_loop___closed__9; +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 = 1; +x_20 = l_Lean_SourceInfo_fromRef(x_9, x_19); +x_21 = l_calc___closed__1; +x_22 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +x_23 = l_calcTactic___closed__2; +lean_inc(x_14); +x_24 = l_Lean_Syntax_node2(x_14, x_23, x_22, x_11); +x_25 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__16; +lean_inc(x_14); +x_26 = l_Lean_Syntax_node1(x_14, x_25, x_24); +x_27 = l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__9; +lean_inc(x_14); +x_28 = l_Lean_Syntax_node1(x_14, x_27, x_26); +x_29 = l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__7; +lean_inc(x_14); +x_30 = l_Lean_Syntax_node1(x_14, x_29, x_28); +x_31 = l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__4; +x_32 = l_Lean_Syntax_node3(x_14, x_31, x_16, x_18, x_30); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_3); +return x_33; +} +} +} static lean_object* _init_l_unexpandUnit___rarg___closed__1() { _start: { @@ -12425,31 +12628,23 @@ static lean_object* _init_l___aux__Init__NotationExtra______macroRules__tacticFu _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Tactic", 6); -return x_1; -} -} -static lean_object* _init_l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__2() { -_start: -{ -lean_object* x_1; x_1 = lean_mk_string_from_bytes("seq1", 4); return x_1; } } -static lean_object* _init_l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__3() { +static lean_object* _init_l___aux__Init__NotationExtra______macroRules__tacticFunext________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_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___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__1; -x_4 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__2; +x_3 = l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__1; +x_4 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___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___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__4() { +static lean_object* _init_l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__3() { _start: { lean_object* x_1; @@ -12457,19 +12652,19 @@ x_1 = lean_mk_string_from_bytes("apply", 5); return x_1; } } -static lean_object* _init_l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__5() { +static lean_object* _init_l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; 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___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__1; -x_4 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__4; +x_3 = l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__1; +x_4 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___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___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__6() { +static lean_object* _init_l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__5() { _start: { lean_object* x_1; lean_object* x_2; @@ -12478,7 +12673,7 @@ x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__7() { +static lean_object* _init_l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -12488,31 +12683,31 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__8() { +static lean_object* _init_l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__7; +x_2 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___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___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__9() { +static lean_object* _init_l___aux__Init__NotationExtra______macroRules__tacticFunext________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___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__8; +x_2 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__7; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__10() { +static lean_object* _init_l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__9() { _start: { lean_object* x_1; @@ -12520,7 +12715,7 @@ x_1 = lean_mk_string_from_bytes(";", 1); return x_1; } } -static lean_object* _init_l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__11() { +static lean_object* _init_l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__10() { _start: { lean_object* x_1; @@ -12528,14 +12723,14 @@ x_1 = lean_mk_string_from_bytes("intro", 5); return x_1; } } -static lean_object* _init_l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__12() { +static lean_object* _init_l___aux__Init__NotationExtra______macroRules__tacticFunext________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___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__1; -x_4 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__11; +x_3 = l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__1; +x_4 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__10; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } @@ -12553,30 +12748,30 @@ lean_inc(x_9); x_10 = lean_ctor_get(x_4, 1); lean_inc(x_10); lean_dec(x_4); -x_11 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__4; +x_11 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___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___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__7; +x_13 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__6; x_14 = l_Lean_addMacroScope(x_10, x_13, x_9); -x_15 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__6; -x_16 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__9; +x_15 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__5; +x_16 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__8; lean_inc(x_8); x_17 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_17, 0, x_8); lean_ctor_set(x_17, 1, x_15); lean_ctor_set(x_17, 2, x_14); lean_ctor_set(x_17, 3, x_16); -x_18 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__5; +x_18 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__4; lean_inc(x_8); x_19 = l_Lean_Syntax_node2(x_8, x_18, x_12, x_17); -x_20 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__10; +x_20 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__9; lean_inc(x_8); x_21 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_21, 0, x_8); lean_ctor_set(x_21, 1, x_20); -x_22 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__11; +x_22 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__10; lean_inc(x_8); x_23 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_23, 0, x_8); @@ -12584,7 +12779,7 @@ lean_ctor_set(x_23, 1, x_22); x_24 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__16; lean_inc(x_8); x_25 = l_Lean_Syntax_node1(x_8, x_24, x_2); -x_26 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__12; +x_26 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__11; lean_inc(x_8); x_27 = l_Lean_Syntax_node2(x_8, x_26, x_23, x_25); x_28 = l_tacticFunext_________closed__3; @@ -12604,7 +12799,7 @@ x_33 = l_Lean_Syntax_node2(x_8, x_1, x_29, x_32); lean_inc(x_21); lean_inc(x_8); x_34 = l_Lean_Syntax_node5(x_8, x_24, x_19, x_21, x_27, x_21, x_33); -x_35 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__3; +x_35 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__2; x_36 = l_Lean_Syntax_node1(x_8, x_35, x_34); x_37 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_37, 0, x_36); @@ -12634,7 +12829,7 @@ static lean_object* _init_l___aux__Init__NotationExtra______macroRules__tacticFu 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___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__1; +x_3 = l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__1; x_4 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__2; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -12651,50 +12846,10 @@ return x_1; static lean_object* _init_l___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__5() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("tacticSeq", 9); -return x_1; -} -} -static lean_object* _init_l___aux__Init__NotationExtra______macroRules__tacticFunext________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_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___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__1; -x_4 = l___aux__Init__NotationExtra______macroRules__tacticFunext________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___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__7() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("tacticSeq1Indented", 18); -return x_1; -} -} -static lean_object* _init_l___aux__Init__NotationExtra______macroRules__tacticFunext________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_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___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__1; -x_4 = l___aux__Init__NotationExtra______macroRules__tacticFunext________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___aux__Init__NotationExtra______macroRules__tacticFunext________1___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_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___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__1; +x_3 = l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__1; x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__17; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -12784,30 +12939,30 @@ lean_inc(x_31); x_32 = lean_ctor_get(x_2, 1); lean_inc(x_32); lean_dec(x_2); -x_33 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__4; +x_33 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__3; lean_inc(x_30); x_34 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_34, 0, x_30); lean_ctor_set(x_34, 1, x_33); -x_35 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__7; +x_35 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__6; x_36 = l_Lean_addMacroScope(x_32, x_35, x_31); -x_37 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__6; -x_38 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__9; +x_37 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__5; +x_38 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__8; lean_inc(x_30); x_39 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_39, 0, x_30); lean_ctor_set(x_39, 1, x_37); lean_ctor_set(x_39, 2, x_36); lean_ctor_set(x_39, 3, x_38); -x_40 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__5; +x_40 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__4; lean_inc(x_30); x_41 = l_Lean_Syntax_node2(x_30, x_40, x_34, x_39); -x_42 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__10; +x_42 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__9; lean_inc(x_30); x_43 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_43, 0, x_30); lean_ctor_set(x_43, 1, x_42); -x_44 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__11; +x_44 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__10; lean_inc(x_30); x_45 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_45, 0, x_30); @@ -12815,12 +12970,12 @@ lean_ctor_set(x_45, 1, x_44); x_46 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__16; lean_inc(x_30); x_47 = l_Lean_Syntax_node1(x_30, x_46, x_27); -x_48 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__12; +x_48 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__11; lean_inc(x_30); x_49 = l_Lean_Syntax_node2(x_30, x_48, x_45, x_47); lean_inc(x_30); x_50 = l_Lean_Syntax_node3(x_30, x_46, x_41, x_43, x_49); -x_51 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__3; +x_51 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__2; x_52 = l_Lean_Syntax_node1(x_30, x_51, x_50); x_53 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_53, 0, x_52); @@ -12851,30 +13006,30 @@ lean_inc(x_56); x_62 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_62, 0, x_56); lean_ctor_set(x_62, 1, x_61); -x_63 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__4; +x_63 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__3; lean_inc(x_56); x_64 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_64, 0, x_56); lean_ctor_set(x_64, 1, x_63); -x_65 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__7; +x_65 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__6; x_66 = l_Lean_addMacroScope(x_58, x_65, x_57); -x_67 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__6; -x_68 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__9; +x_67 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__5; +x_68 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__8; lean_inc(x_56); x_69 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_69, 0, x_56); lean_ctor_set(x_69, 1, x_67); lean_ctor_set(x_69, 2, x_66); lean_ctor_set(x_69, 3, x_68); -x_70 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__5; +x_70 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__4; lean_inc(x_56); x_71 = l_Lean_Syntax_node2(x_56, x_70, x_64, x_69); -x_72 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__10; +x_72 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__9; lean_inc(x_56); x_73 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_73, 0, x_56); lean_ctor_set(x_73, 1, x_72); -x_74 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__11; +x_74 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__10; lean_inc(x_56); x_75 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_75, 0, x_56); @@ -12886,15 +13041,15 @@ x_78 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_78, 0, x_56); lean_ctor_set(x_78, 1, x_76); lean_ctor_set(x_78, 2, x_77); -x_79 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__12; +x_79 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__11; lean_inc(x_56); x_80 = l_Lean_Syntax_node2(x_56, x_79, x_75, x_78); lean_inc(x_56); x_81 = l_Lean_Syntax_node3(x_56, x_76, x_71, x_73, x_80); -x_82 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__8; +x_82 = l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__9; lean_inc(x_56); x_83 = l_Lean_Syntax_node1(x_56, x_82, x_81); -x_84 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__6; +x_84 = l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__7; lean_inc(x_56); x_85 = l_Lean_Syntax_node1(x_56, x_84, x_83); x_86 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__23; @@ -12902,7 +13057,7 @@ lean_inc(x_56); x_87 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_87, 0, x_56); lean_ctor_set(x_87, 1, x_86); -x_88 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__9; +x_88 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__5; lean_inc(x_56); x_89 = l_Lean_Syntax_node3(x_56, x_88, x_62, x_85, x_87); lean_inc(x_56); @@ -14730,7 +14885,7 @@ 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___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__10; +x_31 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__9; lean_inc(x_19); x_32 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_32, 0, x_19); @@ -15051,7 +15206,7 @@ lean_inc(x_42); x_59 = l_Lean_Syntax_node5(x_42, x_16, x_20, x_48, x_56, x_58, x_36); lean_inc(x_42); x_60 = l_Lean_Syntax_node1(x_42, x_10, x_59); -x_61 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__10; +x_61 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__9; lean_inc(x_42); x_62 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_62, 0, x_42); @@ -15177,7 +15332,7 @@ lean_inc(x_85); x_105 = l_Lean_Syntax_node5(x_85, x_16, x_94, x_98, x_102, x_104, x_79); lean_inc(x_85); x_106 = l_Lean_Syntax_node1(x_85, x_10, x_105); -x_107 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__10; +x_107 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__9; lean_inc(x_85); x_108 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_108, 0, x_85); @@ -15260,7 +15415,7 @@ lean_inc(x_118); x_142 = l_Lean_Syntax_node5(x_118, x_16, x_127, x_131, x_139, x_141, x_112); lean_inc(x_118); x_143 = l_Lean_Syntax_node1(x_118, x_10, x_142); -x_144 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__10; +x_144 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__9; lean_inc(x_118); x_145 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_145, 0, x_118); @@ -15415,7 +15570,7 @@ 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___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__10; +x_31 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__9; lean_inc(x_19); x_32 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_32, 0, x_19); @@ -15651,7 +15806,7 @@ lean_inc(x_42); x_59 = l_Lean_Syntax_node5(x_42, x_16, x_20, x_48, x_56, x_58, x_36); lean_inc(x_42); x_60 = l_Lean_Syntax_node1(x_42, x_10, x_59); -x_61 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__10; +x_61 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__9; lean_inc(x_42); x_62 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_62, 0, x_42); @@ -15777,7 +15932,7 @@ lean_inc(x_85); x_105 = l_Lean_Syntax_node5(x_85, x_16, x_94, x_98, x_102, x_104, x_79); lean_inc(x_85); x_106 = l_Lean_Syntax_node1(x_85, x_10, x_105); -x_107 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__10; +x_107 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__9; lean_inc(x_85); x_108 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_108, 0, x_85); @@ -15860,7 +16015,7 @@ lean_inc(x_118); x_142 = l_Lean_Syntax_node5(x_118, x_16, x_127, x_131, x_139, x_141, x_112); lean_inc(x_118); x_143 = l_Lean_Syntax_node1(x_118, x_10, x_142); -x_144 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__10; +x_144 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__9; lean_inc(x_118); x_145 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_145, 0, x_118); @@ -16315,7 +16470,7 @@ static lean_object* _init_l_solve___closed__11() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__5; +x_2 = l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -16510,7 +16665,7 @@ static lean_object* _init_l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra 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___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__1; +x_3 = l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__1; x_4 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__solve__1___spec__3___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -16550,10 +16705,10 @@ lean_inc(x_1); x_17 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_17, 0, x_1); lean_ctor_set(x_17, 1, x_16); -x_18 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__9; +x_18 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__5; lean_inc(x_1); x_19 = l_Lean_Syntax_node3(x_1, x_18, x_15, x_9, x_17); -x_20 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__10; +x_20 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__9; lean_inc(x_1); x_21 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_21, 0, x_1); @@ -16635,7 +16790,7 @@ static lean_object* _init_l___aux__Init__NotationExtra______macroRules__solve__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_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___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__1; +x_3 = l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__1; x_4 = l___aux__Init__NotationExtra______macroRules__solve__1___closed__2; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -16655,7 +16810,7 @@ static lean_object* _init_l___aux__Init__NotationExtra______macroRules__solve__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_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___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__1; +x_3 = l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__1; x_4 = l___aux__Init__NotationExtra______macroRules__solve__1___closed__4; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -16725,8 +16880,8 @@ x_23 = lean_array_get_size(x_15); x_24 = lean_usize_of_nat(x_23); lean_dec(x_23); x_25 = 0; -x_26 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__6; -x_27 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__8; +x_26 = l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__7; +x_27 = l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__9; x_28 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__16; lean_inc(x_18); x_29 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__solve__1___spec__3(x_18, x_26, x_27, x_28, x_24, x_25, x_15); @@ -18032,7 +18187,7 @@ lean_ctor_set(x_18, 1, x_17); 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; +x_21 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__9; lean_inc(x_14); x_22 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_22, 0, x_14); @@ -19308,6 +19463,7 @@ return x_36; lean_object* initialize_Init_Meta(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_Array_Subarray(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_ToString(uint8_t builtin, lean_object*); +lean_object* initialize_Init_Conv(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Init_NotationExtra(uint8_t builtin, lean_object* w) { lean_object * res; @@ -19322,6 +19478,9 @@ lean_dec_ref(res); res = initialize_Init_Data_ToString(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Init_Conv(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); l_Lean_termMacro_x2etrace_x5b___x5d_____closed__1 = _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__1(); lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b___x5d_____closed__1); l_Lean_termMacro_x2etrace_x5b___x5d_____closed__2 = _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__2(); @@ -20036,6 +20195,32 @@ l_calcTactic___closed__5 = _init_l_calcTactic___closed__5(); lean_mark_persistent(l_calcTactic___closed__5); l_calcTactic = _init_l_calcTactic(); lean_mark_persistent(l_calcTactic); +l_convCalc_____closed__1 = _init_l_convCalc_____closed__1(); +lean_mark_persistent(l_convCalc_____closed__1); +l_convCalc_____closed__2 = _init_l_convCalc_____closed__2(); +lean_mark_persistent(l_convCalc_____closed__2); +l_convCalc_____closed__3 = _init_l_convCalc_____closed__3(); +lean_mark_persistent(l_convCalc_____closed__3); +l_convCalc__ = _init_l_convCalc__(); +lean_mark_persistent(l_convCalc__); +l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__1 = _init_l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__1(); +lean_mark_persistent(l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__1); +l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__2 = _init_l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__2(); +lean_mark_persistent(l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__2); +l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__3 = _init_l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__3(); +lean_mark_persistent(l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__3); +l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__4 = _init_l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__4(); +lean_mark_persistent(l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__4); +l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__5 = _init_l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__5(); +lean_mark_persistent(l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__5); +l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__6 = _init_l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__6(); +lean_mark_persistent(l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__6); +l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__7 = _init_l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__7(); +lean_mark_persistent(l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__7); +l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__8 = _init_l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__8(); +lean_mark_persistent(l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__8); +l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__9 = _init_l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__9(); +lean_mark_persistent(l___aux__Init__NotationExtra______macroRules__convCalc____1___closed__9); l_unexpandUnit___rarg___closed__1 = _init_l_unexpandUnit___rarg___closed__1(); lean_mark_persistent(l_unexpandUnit___rarg___closed__1); l_unexpandUnit___rarg___closed__2 = _init_l_unexpandUnit___rarg___closed__2(); @@ -20180,8 +20365,6 @@ l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2_ lean_mark_persistent(l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__10); l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__11 = _init_l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__11(); lean_mark_persistent(l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__11); -l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__12 = _init_l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__12(); -lean_mark_persistent(l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__12); l___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__1 = _init_l___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__1(); lean_mark_persistent(l___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__1); l___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__2 = _init_l___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__2(); @@ -20192,14 +20375,6 @@ l___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__4 lean_mark_persistent(l___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__4); l___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__5 = _init_l___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__5(); lean_mark_persistent(l___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__5); -l___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__6 = _init_l___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__6(); -lean_mark_persistent(l___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__6); -l___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__7 = _init_l___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__7(); -lean_mark_persistent(l___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__7); -l___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__8 = _init_l___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__8(); -lean_mark_persistent(l___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__8); -l___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__9 = _init_l___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__9(); -lean_mark_persistent(l___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__9); l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__1 = _init_l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__1(); lean_mark_persistent(l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__1); l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__2 = _init_l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__2(); diff --git a/stage0/stdlib/Lean/Compiler/CSimpAttr.c b/stage0/stdlib/Lean/Compiler/CSimpAttr.c index 6b8c39ffecd5..7325523591ef 100644 --- a/stage0/stdlib/Lean/Compiler/CSimpAttr.c +++ b/stage0/stdlib/Lean/Compiler/CSimpAttr.c @@ -89,7 +89,6 @@ lean_object* l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContex static lean_object* l_Lean_Compiler_CSimp_State_thmNames___default___closed__1; lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_CSimp_initFn____x40_Lean_Compiler_CSimpAttr___hyg_464____closed__15; -lean_object* l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__4(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Compiler_CSimp_replaceConstants___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_CSimp_add___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_ReplaceImpl_replaceUnsafeM_visit___at_Lean_Compiler_CSimp_replaceConstants___spec__7(lean_object*, lean_object*, lean_object*); @@ -169,6 +168,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_CSimp_State_switch(lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); lean_object* l_Lean_SMap_insert___at_Lean_NameSSet_insert___spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_CSimp_initFn____x40_Lean_Compiler_CSimpAttr___hyg_132____closed__3; +lean_object* l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6351____spec__4(lean_object*); LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Compiler_CSimp_add___spec__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_getConstInfo___at___private_Lean_Compiler_CSimpAttr_0__Lean_Compiler_CSimp_isConstantReplacement_x3f___spec__1___closed__4; lean_object* lean_nat_add(lean_object*, lean_object*); @@ -372,7 +372,7 @@ lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_3 = lean_ctor_get(x_1, 0); x_4 = lean_ctor_get(x_1, 1); x_5 = l_Lean_SMap_switch___at_Lean_Compiler_CSimp_State_switch___spec__1(x_3); -x_6 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__4(x_4); +x_6 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6351____spec__4(x_4); lean_ctor_set(x_1, 1, x_6); lean_ctor_set(x_1, 0, x_5); return x_1; @@ -386,7 +386,7 @@ lean_inc(x_8); lean_inc(x_7); lean_dec(x_1); x_9 = l_Lean_SMap_switch___at_Lean_Compiler_CSimp_State_switch___spec__1(x_7); -x_10 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__4(x_8); +x_10 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6351____spec__4(x_8); x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_9); lean_ctor_set(x_11, 1, x_10); diff --git a/stage0/stdlib/Lean/Compiler/LCNF/FloatLetIn.c b/stage0/stdlib/Lean/Compiler/LCNF/FloatLetIn.c index c6b5ff94b5bb..5b195b38bc75 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/FloatLetIn.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/FloatLetIn.c @@ -8516,7 +8516,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_FloatLetIn_dontFloat___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_Compiler_LCNF_FloatLetIn_dontFloat___closed__1; x_2 = l_Lean_Compiler_LCNF_FloatLetIn_dontFloat___closed__2; -x_3 = lean_unsigned_to_nat(184u); +x_3 = lean_unsigned_to_nat(210u); x_4 = lean_unsigned_to_nat(14u); x_5 = l_Lean_Compiler_LCNF_FloatLetIn_dontFloat___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/LCNF/JoinPoints.c b/stage0/stdlib/Lean/Compiler/LCNF/JoinPoints.c index 6abffb05a0a9..1ec47fda1ee0 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/JoinPoints.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/JoinPoints.c @@ -8448,7 +8448,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_JoinPointContextExtender_replaceF 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_JoinPointContextExtender_replaceFVar___closed__1; x_2 = l_Lean_Compiler_LCNF_JoinPointContextExtender_replaceFVar___closed__2; -x_3 = lean_unsigned_to_nat(184u); +x_3 = lean_unsigned_to_nat(210u); x_4 = lean_unsigned_to_nat(14u); x_5 = l_Lean_Compiler_LCNF_JoinPointContextExtender_replaceFVar___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/Data/HashMap.c b/stage0/stdlib/Lean/Data/HashMap.c index e4559f16b334..20a8f2393823 100644 --- a/stage0/stdlib/Lean/Data/HashMap.c +++ b/stage0/stdlib/Lean/Data/HashMap.c @@ -188,6 +188,7 @@ lean_object* l_Lean_AssocList_replace___rarg(lean_object*, lean_object*, lean_ob lean_object* lean_array_uget(lean_object*, size_t); LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_HashMap_toArray___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_HashMap_toArray___spec__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_HashMapImp_insertIfNew(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Array_groupByKey___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_HashMapImp_forBucketsM___spec__1___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_HashMapImp_foldM___spec__1___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, size_t, lean_object*); @@ -211,6 +212,8 @@ LEAN_EXPORT lean_object* l_Lean_HashMap_numBuckets___boxed(lean_object*, lean_ob lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMapImp_moveEntries___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMapImp_foldBucketsM(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_HashMapImp_insertIfNew___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_HashMap_insertIfNew___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMapImp_foldM(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_HashMapImp_forM___spec__1___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -218,6 +221,7 @@ LEAN_EXPORT lean_object* l_Lean_HashMap_isEmpty(lean_object*, lean_object*, lean LEAN_EXPORT lean_object* l_Lean_HashMap_find_x21___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_AssocList_findEntry_x3f___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMap_instInhabitedHashMap___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_HashMap_insertIfNew(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMapImp_foldBucketsM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMap_instGetElemHashMapOptionTrue(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_HashMap_0__Lean_numBucketsForCapacity(lean_object*); @@ -1592,6 +1596,157 @@ x_3 = lean_alloc_closure((void*)(l_Lean_HashMapImp_insert___rarg), 5, 0); return x_3; } } +LEAN_EXPORT lean_object* l_Lean_HashMapImp_insertIfNew___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint64_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; +x_6 = lean_ctor_get(x_3, 0); +lean_inc(x_6); +x_7 = lean_ctor_get(x_3, 1); +lean_inc(x_7); +x_8 = lean_array_get_size(x_7); +lean_inc(x_2); +lean_inc(x_4); +x_9 = lean_apply_1(x_2, x_4); +x_10 = lean_unbox_uint64(x_9); +lean_dec(x_9); +lean_inc(x_8); +x_11 = lean_hashmap_mk_idx(x_8, x_10); +x_12 = lean_array_uget(x_7, x_11); +lean_inc(x_12); +lean_inc(x_4); +x_13 = l_Lean_AssocList_find_x3f___rarg(x_1, x_4, x_12); +if (lean_obj_tag(x_13) == 0) +{ +uint8_t x_14; +x_14 = !lean_is_exclusive(x_3); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_15 = lean_ctor_get(x_3, 1); +lean_dec(x_15); +x_16 = lean_ctor_get(x_3, 0); +lean_dec(x_16); +x_17 = lean_unsigned_to_nat(1u); +x_18 = lean_nat_add(x_6, x_17); +lean_dec(x_6); +x_19 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_19, 0, x_4); +lean_ctor_set(x_19, 1, x_5); +lean_ctor_set(x_19, 2, x_12); +x_20 = lean_array_uset(x_7, x_11, x_19); +x_21 = l___private_Lean_Data_HashMap_0__Lean_numBucketsForCapacity(x_18); +x_22 = lean_nat_dec_le(x_21, x_8); +lean_dec(x_8); +lean_dec(x_21); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +lean_free_object(x_3); +x_23 = l_Lean_HashMapImp_expand___rarg(x_2, x_18, x_20); +x_24 = lean_box(0); +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 +{ +lean_object* x_26; lean_object* x_27; +lean_dec(x_2); +lean_ctor_set(x_3, 1, x_20); +lean_ctor_set(x_3, 0, x_18); +x_26 = lean_box(0); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_3); +lean_ctor_set(x_27, 1, x_26); +return x_27; +} +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +lean_dec(x_3); +x_28 = lean_unsigned_to_nat(1u); +x_29 = lean_nat_add(x_6, x_28); +lean_dec(x_6); +x_30 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_30, 0, x_4); +lean_ctor_set(x_30, 1, x_5); +lean_ctor_set(x_30, 2, x_12); +x_31 = lean_array_uset(x_7, x_11, x_30); +x_32 = l___private_Lean_Data_HashMap_0__Lean_numBucketsForCapacity(x_29); +x_33 = lean_nat_dec_le(x_32, x_8); +lean_dec(x_8); +lean_dec(x_32); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = l_Lean_HashMapImp_expand___rarg(x_2, x_29, x_31); +x_35 = lean_box(0); +x_36 = lean_alloc_ctor(0, 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_object* x_38; lean_object* x_39; +lean_dec(x_2); +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_29); +lean_ctor_set(x_37, 1, x_31); +x_38 = lean_box(0); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +return x_39; +} +} +} +else +{ +uint8_t x_40; +lean_dec(x_12); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_40 = !lean_is_exclusive(x_13); +if (x_40 == 0) +{ +lean_object* x_41; +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_3); +lean_ctor_set(x_41, 1, x_13); +return x_41; +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_13, 0); +lean_inc(x_42); +lean_dec(x_13); +x_43 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_43, 0, x_42); +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_3); +lean_ctor_set(x_44, 1, x_43); +return x_44; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_HashMapImp_insertIfNew(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Lean_HashMapImp_insertIfNew___rarg), 5, 0); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_HashMapImp_erase___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { @@ -2076,6 +2231,157 @@ x_3 = lean_alloc_closure((void*)(l_Lean_HashMap_insert_x27___rarg), 5, 0); return x_3; } } +LEAN_EXPORT lean_object* l_Lean_HashMap_insertIfNew___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint64_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; +x_6 = lean_ctor_get(x_3, 0); +lean_inc(x_6); +x_7 = lean_ctor_get(x_3, 1); +lean_inc(x_7); +x_8 = lean_array_get_size(x_7); +lean_inc(x_2); +lean_inc(x_4); +x_9 = lean_apply_1(x_2, x_4); +x_10 = lean_unbox_uint64(x_9); +lean_dec(x_9); +lean_inc(x_8); +x_11 = lean_hashmap_mk_idx(x_8, x_10); +x_12 = lean_array_uget(x_7, x_11); +lean_inc(x_12); +lean_inc(x_4); +x_13 = l_Lean_AssocList_find_x3f___rarg(x_1, x_4, x_12); +if (lean_obj_tag(x_13) == 0) +{ +uint8_t x_14; +x_14 = !lean_is_exclusive(x_3); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_15 = lean_ctor_get(x_3, 1); +lean_dec(x_15); +x_16 = lean_ctor_get(x_3, 0); +lean_dec(x_16); +x_17 = lean_unsigned_to_nat(1u); +x_18 = lean_nat_add(x_6, x_17); +lean_dec(x_6); +x_19 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_19, 0, x_4); +lean_ctor_set(x_19, 1, x_5); +lean_ctor_set(x_19, 2, x_12); +x_20 = lean_array_uset(x_7, x_11, x_19); +x_21 = l___private_Lean_Data_HashMap_0__Lean_numBucketsForCapacity(x_18); +x_22 = lean_nat_dec_le(x_21, x_8); +lean_dec(x_8); +lean_dec(x_21); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +lean_free_object(x_3); +x_23 = l_Lean_HashMapImp_expand___rarg(x_2, x_18, x_20); +x_24 = lean_box(0); +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 +{ +lean_object* x_26; lean_object* x_27; +lean_dec(x_2); +lean_ctor_set(x_3, 1, x_20); +lean_ctor_set(x_3, 0, x_18); +x_26 = lean_box(0); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_3); +lean_ctor_set(x_27, 1, x_26); +return x_27; +} +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +lean_dec(x_3); +x_28 = lean_unsigned_to_nat(1u); +x_29 = lean_nat_add(x_6, x_28); +lean_dec(x_6); +x_30 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_30, 0, x_4); +lean_ctor_set(x_30, 1, x_5); +lean_ctor_set(x_30, 2, x_12); +x_31 = lean_array_uset(x_7, x_11, x_30); +x_32 = l___private_Lean_Data_HashMap_0__Lean_numBucketsForCapacity(x_29); +x_33 = lean_nat_dec_le(x_32, x_8); +lean_dec(x_8); +lean_dec(x_32); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = l_Lean_HashMapImp_expand___rarg(x_2, x_29, x_31); +x_35 = lean_box(0); +x_36 = lean_alloc_ctor(0, 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_object* x_38; lean_object* x_39; +lean_dec(x_2); +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_29); +lean_ctor_set(x_37, 1, x_31); +x_38 = lean_box(0); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +return x_39; +} +} +} +else +{ +uint8_t x_40; +lean_dec(x_12); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_40 = !lean_is_exclusive(x_13); +if (x_40 == 0) +{ +lean_object* x_41; +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_3); +lean_ctor_set(x_41, 1, x_13); +return x_41; +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_13, 0); +lean_inc(x_42); +lean_dec(x_13); +x_43 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_43, 0, x_42); +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_3); +lean_ctor_set(x_44, 1, x_43); +return x_44; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_HashMap_insertIfNew(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Lean_HashMap_insertIfNew___rarg), 5, 0); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_HashMap_erase___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { @@ -2210,7 +2516,7 @@ static lean_object* _init_l_Lean_HashMap_find_x21___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_HashMap_find_x21___rarg___closed__1; x_2 = l_Lean_HashMap_find_x21___rarg___closed__2; -x_3 = lean_unsigned_to_nat(184u); +x_3 = lean_unsigned_to_nat(210u); x_4 = lean_unsigned_to_nat(14u); x_5 = l_Lean_HashMap_find_x21___rarg___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/Elab/AuxDef.c b/stage0/stdlib/Lean/Elab/AuxDef.c index c03703f8ad94..05efb969384f 100644 --- a/stage0/stdlib/Lean/Elab/AuxDef.c +++ b/stage0/stdlib/Lean/Elab/AuxDef.c @@ -49,6 +49,7 @@ LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Elab_Command_elabAuxDef___s 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*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabAuxDef___closed__3; +static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__26; static lean_object* l___regBuiltin_Lean_Elab_Command_elabAuxDef_declRange___closed__1; static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__22; static lean_object* l_Lean_Elab_Command_aux__def___closed__19; @@ -870,7 +871,7 @@ static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__1 _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("def", 3); +x_1 = lean_mk_string_from_bytes("definition", 10); return x_1; } } @@ -890,23 +891,31 @@ static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__1 _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("declId", 6); +x_1 = lean_mk_string_from_bytes("def", 3); return x_1; } } static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__15() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("declId", 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___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_Elab_Command_aux__def___closed__1; x_2 = l_Lean_Elab_Command_aux__def___closed__14; x_3 = l_Lean_Elab_Command_aux__def___closed__3; -x_4 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__14; +x_4 = l_Lean_Elab_Command_elabAuxDef___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_Elab_Command_elabAuxDef___lambda__1___closed__16() { +static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__17() { _start: { lean_object* x_1; @@ -914,19 +923,19 @@ x_1 = lean_mk_string_from_bytes("optDeclSig", 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__17() { +static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_Elab_Command_aux__def___closed__1; x_2 = l_Lean_Elab_Command_aux__def___closed__14; x_3 = l_Lean_Elab_Command_aux__def___closed__3; -x_4 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__16; +x_4 = l_Lean_Elab_Command_elabAuxDef___lambda__1___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_Elab_Command_elabAuxDef___lambda__1___closed__18() { +static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__19() { _start: { lean_object* x_1; @@ -934,19 +943,19 @@ x_1 = lean_mk_string_from_bytes("typeSpec", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__19() { +static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__20() { _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_aux__def___closed__1; x_2 = l_Lean_Elab_Command_aux__def___closed__14; x_3 = l_Lean_Elab_Command_aux__def___closed__15; -x_4 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__18; +x_4 = l_Lean_Elab_Command_elabAuxDef___lambda__1___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_Elab_Command_elabAuxDef___lambda__1___closed__20() { +static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__21() { _start: { lean_object* x_1; @@ -954,19 +963,19 @@ x_1 = lean_mk_string_from_bytes("declValSimple", 13); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__21() { +static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__22() { _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_aux__def___closed__1; x_2 = l_Lean_Elab_Command_aux__def___closed__14; x_3 = l_Lean_Elab_Command_aux__def___closed__3; -x_4 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__20; +x_4 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__21; 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_elabAuxDef___lambda__1___closed__22() { +static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__23() { _start: { lean_object* x_1; @@ -974,7 +983,7 @@ x_1 = lean_mk_string_from_bytes("Termination", 11); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__23() { +static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__24() { _start: { lean_object* x_1; @@ -982,19 +991,19 @@ x_1 = lean_mk_string_from_bytes("suffix", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__24() { +static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__25() { _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_aux__def___closed__1; x_2 = l_Lean_Elab_Command_aux__def___closed__14; -x_3 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__22; -x_4 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__23; +x_3 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__23; +x_4 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__24; 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_elabAuxDef___lambda__1___closed__25() { +static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__26() { _start: { lean_object* x_1; lean_object* x_2; @@ -1137,7 +1146,7 @@ x_50 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_50, 0, x_43); lean_ctor_set(x_50, 1, x_48); lean_ctor_set(x_50, 2, x_49); -x_51 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__12; +x_51 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__14; lean_inc(x_43); x_52 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_52, 0, x_43); @@ -1149,7 +1158,7 @@ lean_ctor_set(x_54, 1, x_48); lean_ctor_set(x_54, 2, x_14); x_55 = 1; x_56 = l_Lean_mkIdentFrom(x_54, x_38, x_55); -x_57 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__15; +x_57 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__16; lean_inc(x_50); lean_inc(x_43); x_58 = l_Lean_Syntax_node2(x_43, x_57, x_56, x_50); @@ -1158,12 +1167,12 @@ lean_inc(x_43); x_60 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_60, 0, x_43); lean_ctor_set(x_60, 1, x_59); -x_61 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__19; +x_61 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__20; lean_inc(x_43); x_62 = l_Lean_Syntax_node2(x_43, x_61, x_60, x_11); lean_inc(x_43); x_63 = l_Lean_Syntax_node1(x_43, x_48, x_62); -x_64 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__17; +x_64 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__18; lean_inc(x_50); lean_inc(x_43); x_65 = l_Lean_Syntax_node2(x_43, x_64, x_50, x_63); @@ -1172,11 +1181,11 @@ lean_inc(x_43); x_67 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_67, 0, x_43); lean_ctor_set(x_67, 1, x_66); -x_68 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__24; +x_68 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__25; lean_inc_n(x_50, 2); lean_inc(x_43); x_69 = l_Lean_Syntax_node2(x_43, x_68, x_50, x_50); -x_70 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__21; +x_70 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__22; lean_inc(x_50); lean_inc(x_43); x_71 = l_Lean_Syntax_node4(x_43, x_70, x_67, x_13, x_69, x_50); @@ -1187,7 +1196,7 @@ x_73 = l_Lean_Syntax_node5(x_43, x_72, x_52, x_58, x_65, x_71, x_50); if (lean_obj_tag(x_2) == 0) { lean_object* x_74; lean_object* x_75; -x_74 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__25; +x_74 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__26; lean_inc(x_43); x_75 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_75, 0, x_43); @@ -1245,7 +1254,7 @@ lean_ctor_set(x_93, 2, x_92); if (lean_obj_tag(x_4) == 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; -x_94 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__25; +x_94 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__26; lean_inc(x_43); x_95 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_95, 0, x_43); @@ -1797,6 +1806,8 @@ l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__24 = _init_l_Lean_Elab_Comm lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__24); l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__25 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__25(); lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__25); +l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__26 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__26(); +lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__26); l_Lean_Elab_Command_elabAuxDef___closed__1 = _init_l_Lean_Elab_Command_elabAuxDef___closed__1(); lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___closed__1); l___regBuiltin_Lean_Elab_Command_elabAuxDef___closed__1 = _init_l___regBuiltin_Lean_Elab_Command_elabAuxDef___closed__1(); diff --git a/stage0/stdlib/Lean/Elab/Declaration.c b/stage0/stdlib/Lean/Elab/Declaration.c index 966133c93bba..be610870e98b 100644 --- a/stage0/stdlib/Lean/Elab/Declaration.c +++ b/stage0/stdlib/Lean/Elab/Declaration.c @@ -20,6 +20,7 @@ static lean_object* l_Lean_Elab_Command_expandInitialize___lambda__1___closed__7 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_Command_expandMutualNamespace(lean_object*); static lean_object* l_Lean_Elab_Command_expandInitialize___lambda__1___closed__20; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__11; static lean_object* l_Lean_Elab_Command_expandInitialize___lambda__1___closed__14; LEAN_EXPORT lean_object* l_Lean_Elab_Command_checkValidCtorModifier___at___private_Lean_Elab_Declaration_0__Lean_Elab_Command_inductiveSyntaxToView___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_setDefName___closed__3; @@ -42,6 +43,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Command_elabMutual_declRange___clos static lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_ensureValidNamespace___closed__3; lean_object* l_Lean_addDocString_x27___at_Lean_Elab_Command_expandDeclId___spec__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabDeclaration___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__12; static lean_object* l_Lean_Elab_Command_expandInitialize___lambda__1___closed__45; LEAN_EXPORT lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_expandDeclNamespace_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_expandMutualNamespace_declRange(lean_object*); @@ -80,6 +82,7 @@ LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_Command_expandMutualNamespace___ lean_object* l_Lean_Name_toString(lean_object*, uint8_t); LEAN_EXPORT lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_expandDeclNamespace_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getId(lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__7; static lean_object* l_Lean_Elab_Command_checkValidCtorModifier___at___private_Lean_Elab_Declaration_0__Lean_Elab_Command_inductiveSyntaxToView___spec__1___lambda__1___closed__2; static lean_object* l_Lean_Elab_Command_elabAxiom___lambda__5___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Command_expandMutualNamespace_declRange___closed__4; @@ -129,7 +132,6 @@ static lean_object* l_Lean_Elab_Command_elabDeclaration___closed__1; static lean_object* l_Lean_Elab_Command_expandInitialize___lambda__1___closed__12; static lean_object* l_Lean_Elab_Command_elabAxiom___lambda__5___closed__7; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Declaration_0__Lean_Elab_Command_inductiveSyntaxToView___spec__3___lambda__3(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_Declaration___hyg_7880____closed__5; LEAN_EXPORT lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_elabMutualInductive(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Command_elabAttr___spec__3___closed__2; static lean_object* l_Lean_Elab_Command_expandInitialize___lambda__1___closed__29; @@ -144,6 +146,7 @@ lean_object* l_Lean_Elab_elabModifiers___at_Lean_Elab_Command_elabMutualDef___sp static lean_object* l___regBuiltin_Lean_Elab_Command_elabDeclaration___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandInitialize___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isMutualPreambleCommand(lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__4; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabAttr___spec__4___lambda__1(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_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*); @@ -153,8 +156,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabDeclaration_declRange(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_expandMutualElement___closed__2; static lean_object* l_Lean_Elab_Command_expandInitialize___lambda__1___closed__3; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__4; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__11; 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___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__20; @@ -170,6 +171,7 @@ lean_object* l_Lean_CollectLevelParams_main(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__6; lean_object* l_Lean_Elab_Command_checkValidInductiveModifier___at_Lean_Elab_Command_elabStructure___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_expandDeclNamespace_x3f(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__23; lean_object* l_Lean_Elab_withSaveInfoContext___at___private_Lean_Elab_PreDefinition_Basic_0__Lean_Elab_addNonRecAux___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_checkValidCtorModifier___at___private_Lean_Elab_Declaration_0__Lean_Elab_Command_inductiveSyntaxToView___spec__1___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_expandInitialize_declRange(lean_object*); @@ -179,7 +181,6 @@ static lean_object* l_Lean_Elab_Command_checkValidCtorModifier___at___private_Le LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabAttr_declRange(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_expandInitialize(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabDeclaration_declRange___closed__1; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__8; static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabDeclaration___spec__4___closed__1; lean_object* l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSavedPartialInfoContext___at___private_Lean_Elab_PreDefinition_Basic_0__Lean_Elab_addNonRecAux___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_expandMutualPreamble___closed__3; @@ -204,7 +205,6 @@ uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_expandInitialize___closed__1; lean_object* l_Lean_Elab_expandMacroImpl_x3f(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_expandMutualNamespace___closed__2; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__2; lean_object* l_Lean_Elab_getDeclarationSelectionRef(lean_object*); lean_object* l_Lean_ResolveName_resolveNamespace(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__13; @@ -214,7 +214,7 @@ static lean_object* l_Lean_Elab_Command_elabAttr___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandInitialize___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_expandInitialize___lambda__1___closed__39; lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__9; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__13; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabDeclaration___spec__5(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_expandInitialize___lambda__1___closed__24; uint8_t l_Lean_Elab_Modifiers_isProtected(lean_object*); @@ -225,6 +225,8 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_expandMut LEAN_EXPORT uint8_t l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef(lean_object*); static lean_object* l_Lean_Elab_Command_expandInitialize___lambda__1___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Command_expandMutualNamespace_declRange___closed__2; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__2; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__6; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabAttr___spec__1___closed__4; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Declaration_0__Lean_Elab_Command_elabMutualInductive___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); 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*); @@ -250,7 +252,6 @@ static lean_object* l_Lean_Elab_Command_checkValidCtorModifier___at___private_Le lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_expandInitialize___lambda__1___closed__36; static lean_object* l_Lean_Elab_Command_elabAxiom___lambda__5___closed__3; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__13; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Declaration_0__Lean_Elab_Command_inductiveSyntaxToView___spec__3___closed__2; lean_object* l_Lean_Elab_elabAttrs___at_Lean_Elab_Command_elabMutualDef___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Command_liftTermElabM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -285,7 +286,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_getDeclarationRange___at_Lean_Elab_Command_ static lean_object* l_Lean_Elab_Command_expandInitialize___lambda__5___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_getDeclarationRange___at_Lean_Elab_Command_elabAxiom___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_expandInitialize___lambda__1___closed__34; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__7; static lean_object* l_Lean_Elab_Command_expandInitialize___lambda__2___closed__3; static lean_object* l_Lean_Elab_Command_expandInitialize___lambda__1___closed__26; static lean_object* l___regBuiltin_Lean_Elab_Command_expandMutualElement_declRange___closed__5; @@ -301,8 +301,6 @@ static lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_induc static lean_object* l___regBuiltin_Lean_Elab_Command_elabDeclaration_declRange___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Command_expandMutualPreamble_declRange___closed__6; lean_object* l_Lean_Elab_Command_getCurrMacroScope(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__1; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__12; static lean_object* l___regBuiltin_Lean_Elab_Command_expandMutualPreamble_declRange___closed__7; static lean_object* l_Lean_Elab_Command_expandInitialize___lambda__1___closed__16; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Declaration_0__Lean_Elab_Command_inductiveSyntaxToView___spec__3(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); @@ -339,6 +337,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabDeclaration(lean_o uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabDeclaration___spec__5___rarg___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandMutualNamespace___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__22; static lean_object* l___regBuiltin_Lean_Elab_Command_elabAttr_declRange___closed__1; uint8_t l_Lean_Syntax_isIdent(lean_object*); static lean_object* l_Lean_Elab_Command_elabMutual___closed__2; @@ -364,7 +363,6 @@ static lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isMut 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*); -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880_(lean_object*); lean_object* l_Lean_MessageData_ofExpr(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_expandMutualElement_declRange___closed__7; LEAN_EXPORT lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_splitMutualPreamble_loop(lean_object*, lean_object*); @@ -374,12 +372,14 @@ static lean_object* l___regBuiltin_Lean_Elab_Command_expandMutualElement___close static lean_object* l_Lean_Elab_Command_expandInitialize___lambda__1___closed__35; lean_object* l_panic___at_Lean_Parser_SyntaxStack_back___spec__1(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabDeclaration___closed__3; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__8; lean_object* l_Lean_Syntax_node4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabMutual___closed__2; static lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_setDeclIdName___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Command_elabAttr___closed__3; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Declaration_0__Lean_Elab_Command_elabMutualInductive___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__16; +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888_(lean_object*); lean_object* l_Lean_Syntax_setArg(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Elab_Modifiers_isPrivate(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Command_elabAttr___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -387,6 +387,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_ static lean_object* l___regBuiltin_Lean_Elab_Command_elabMutual_declRange___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Command_elabDeclaration_declRange___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Command_elabMutual_declRange___closed__2; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__3; 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_FileMap_leanPosToLspPos(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_expandInitialize___lambda__1___closed__23; @@ -400,6 +401,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_ static lean_object* l___regBuiltin_Lean_Elab_Command_elabAttr___closed__5; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Declaration_0__Lean_Elab_Command_inductiveSyntaxToView___spec__3___lambda__3___closed__2; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isMutualDef___spec__1(lean_object*, size_t, size_t); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__5; static lean_object* l_Lean_Elab_Command_checkValidCtorModifier___at___private_Lean_Elab_Declaration_0__Lean_Elab_Command_inductiveSyntaxToView___spec__1___lambda__2___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Command_expandMutualPreamble_declRange___closed__2; static lean_object* l_Lean_Elab_Command_expandInitialize___lambda__1___closed__38; @@ -411,7 +413,6 @@ static lean_object* l_Lean_Elab_Command_elabClassInductive___closed__1; static lean_object* l_Lean_Elab_Command_elabDeclaration___closed__8; LEAN_EXPORT lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_ensureValidNamespace(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabDeclaration(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__3; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isMutualDef___spec__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Command_elabAxiom___spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandMutualNamespace(lean_object*, lean_object*, lean_object*); @@ -432,6 +433,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_checkValidCtorModifier___at___priva static lean_object* l___regBuiltin_Lean_Elab_Command_expandMutualPreamble___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Command_expandInitialize_declRange___closed__1; lean_object* l_Lean_Elab_Command_elabInductiveViews(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__9; LEAN_EXPORT lean_object* l_Lean_addDeclarationRanges___at_Lean_Elab_Command_elabAxiom___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_setDeclIdName___closed__6; uint8_t l_Lean_isAttribute(lean_object*, lean_object*); @@ -449,6 +451,7 @@ static lean_object* l_Lean_Elab_Command_elabAxiom___lambda__4___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Command_expandInitialize_declRange___closed__5; static lean_object* l_Lean_Elab_Command_expandInitialize___lambda__1___closed__5; static lean_object* l_Lean_Elab_Command_expandMutualPreamble___closed__2; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__10; LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandMutualNamespace___lambda__1(lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabDeclaration___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__5; @@ -462,7 +465,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_checkValidCtorModifier___at___priva lean_object* l_Lean_Syntax_setInfo(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabAttr___spec__4(lean_object*, size_t, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ResolveName_resolveGlobalName(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__10; lean_object* l_Array_mkArray1___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandInitialize___lambda__5(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_expandInitialize___closed__3; @@ -489,6 +491,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Command_elabMutual_declRange___clos static lean_object* l_Lean_Elab_Command_elabDeclaration___closed__5; size_t lean_usize_add(size_t, size_t); static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabDeclaration___spec__5___rarg___closed__2; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__1; static lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandInitialize___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_Elab_Command_expandDeclId(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -520,7 +523,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_ static lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isMutualPreambleCommand___closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabAxiom___lambda__1___boxed(lean_object*); static lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_setDefName___closed__4; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandMutualElement(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Declaration_0__Lean_Elab_Command_inductiveSyntaxToView___spec__4(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_setDefName___closed__2; @@ -901,7 +903,7 @@ static lean_object* _init_l___private_Lean_Elab_Declaration_0__Lean_Elab_Command _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("def", 3); +x_1 = lean_mk_string_from_bytes("definition", 10); return x_1; } } @@ -921,7 +923,7 @@ static lean_object* _init_l___private_Lean_Elab_Declaration_0__Lean_Elab_Command _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("theorem", 7); +x_1 = lean_mk_string_from_bytes("def", 3); return x_1; } } @@ -941,7 +943,7 @@ static lean_object* _init_l___private_Lean_Elab_Declaration_0__Lean_Elab_Command _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("opaque", 6); +x_1 = lean_mk_string_from_bytes("theorem", 7); return x_1; } } @@ -961,7 +963,7 @@ static lean_object* _init_l___private_Lean_Elab_Declaration_0__Lean_Elab_Command _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("axiom", 5); +x_1 = lean_mk_string_from_bytes("opaque", 6); return x_1; } } @@ -981,7 +983,7 @@ static lean_object* _init_l___private_Lean_Elab_Declaration_0__Lean_Elab_Command _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("inductive", 9); +x_1 = lean_mk_string_from_bytes("axiom", 5); return x_1; } } @@ -1001,7 +1003,7 @@ static lean_object* _init_l___private_Lean_Elab_Declaration_0__Lean_Elab_Command _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("classInductive", 14); +x_1 = lean_mk_string_from_bytes("inductive", 9); return x_1; } } @@ -1021,7 +1023,7 @@ static lean_object* _init_l___private_Lean_Elab_Declaration_0__Lean_Elab_Command _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("structure", 9); +x_1 = lean_mk_string_from_bytes("classInductive", 14); return x_1; } } @@ -1037,6 +1039,26 @@ x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } +static lean_object* _init_l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__22() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("structure", 9); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___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___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__1; +x_2 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__2; +x_3 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__3; +x_4 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__22; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} LEAN_EXPORT uint8_t l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef(lean_object* x_1) { _start: { @@ -1095,25 +1117,14 @@ if (x_21 == 0) lean_object* x_22; uint8_t x_23; x_22 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__21; x_23 = lean_name_eq(x_7, x_22); -lean_dec(x_7); -return x_23; -} -else -{ -uint8_t x_24; -lean_dec(x_7); -x_24 = 1; -return x_24; -} -} -else +if (x_23 == 0) { -uint8_t x_25; +lean_object* x_24; uint8_t x_25; +x_24 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__23; +x_25 = lean_name_eq(x_7, x_24); lean_dec(x_7); -x_25 = 1; return x_25; } -} else { uint8_t x_26; @@ -1154,6 +1165,30 @@ x_30 = 1; return x_30; } } +else +{ +uint8_t x_31; +lean_dec(x_7); +x_31 = 1; +return x_31; +} +} +else +{ +uint8_t x_32; +lean_dec(x_7); +x_32 = 1; +return x_32; +} +} +else +{ +uint8_t x_33; +lean_dec(x_7); +x_33 = 1; +return x_33; +} +} } } LEAN_EXPORT lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___boxed(lean_object* x_1) { @@ -1325,7 +1360,7 @@ static lean_object* _init_l___private_Lean_Elab_Declaration_0__Lean_Elab_Command 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_Declaration_0__Lean_Elab_Command_setDeclIdName___closed__5; x_2 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_setDefName___closed__3; -x_3 = lean_unsigned_to_nat(81u); +x_3 = lean_unsigned_to_nat(82u); x_4 = lean_unsigned_to_nat(4u); x_5 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_setDefName___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -2244,7 +2279,7 @@ static lean_object* _init_l_Lean_Elab_Command_elabAxiom___lambda__5___closed__6( { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Command_elabAxiom___lambda__5___closed__5; -x_2 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__14; +x_2 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__16; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -5177,22 +5212,22 @@ x_9 = lean_unsigned_to_nat(1u); x_10 = l_Lean_Syntax_getArg(x_1, x_9); lean_inc(x_10); x_11 = l_Lean_Syntax_getKind(x_10); -x_12 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__15; +x_12 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__17; x_13 = lean_name_eq(x_11, x_12); if (x_13 == 0) { lean_object* x_14; uint8_t x_15; -x_14 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__17; +x_14 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__19; x_15 = lean_name_eq(x_11, x_14); if (x_15 == 0) { lean_object* x_16; uint8_t x_17; -x_16 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__19; +x_16 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__21; x_17 = lean_name_eq(x_11, x_16); if (x_17 == 0) { lean_object* x_18; uint8_t x_19; -x_18 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__21; +x_18 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__23; x_19 = lean_name_eq(x_11, x_18); lean_dec(x_11); if (x_19 == 0) @@ -5612,7 +5647,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabDeclaration_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(192u); +x_1 = lean_unsigned_to_nat(193u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -5624,7 +5659,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabDeclaration_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(216u); +x_1 = lean_unsigned_to_nat(217u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -5652,7 +5687,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabDeclaration_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(192u); +x_1 = lean_unsigned_to_nat(193u); x_2 = lean_unsigned_to_nat(4u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -5664,7 +5699,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabDeclaration_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(192u); +x_1 = lean_unsigned_to_nat(193u); x_2 = lean_unsigned_to_nat(19u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -5723,7 +5758,7 @@ x_6 = lean_unsigned_to_nat(1u); x_7 = l_Lean_Syntax_getArg(x_5, x_6); lean_dec(x_5); x_8 = l_Lean_Syntax_getKind(x_7); -x_9 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__17; +x_9 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__19; x_10 = lean_name_eq(x_8, x_9); lean_dec(x_8); if (x_10 == 0) @@ -6608,7 +6643,7 @@ static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_expandMu 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_Declaration_0__Lean_Elab_Command_setDeclIdName___closed__5; x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_expandMutualNamespace___spec__3___closed__1; -x_3 = lean_unsigned_to_nat(294u); +x_3 = lean_unsigned_to_nat(295u); x_4 = lean_unsigned_to_nat(40u); x_5 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_expandMutualNamespace___spec__3___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -7089,7 +7124,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandMutualNamespace _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(285u); +x_1 = lean_unsigned_to_nat(286u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7101,7 +7136,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandMutualNamespace _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(300u); +x_1 = lean_unsigned_to_nat(301u); x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7129,7 +7164,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandMutualNamespace _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(285u); +x_1 = lean_unsigned_to_nat(286u); x_2 = lean_unsigned_to_nat(4u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7141,7 +7176,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandMutualNamespace _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(285u); +x_1 = lean_unsigned_to_nat(286u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7567,7 +7602,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandMutualElement_d _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(303u); +x_1 = lean_unsigned_to_nat(304u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7579,7 +7614,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandMutualElement_d _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(313u); +x_1 = lean_unsigned_to_nat(314u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7607,7 +7642,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandMutualElement_d _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(303u); +x_1 = lean_unsigned_to_nat(304u); x_2 = lean_unsigned_to_nat(4u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7619,7 +7654,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandMutualElement_d _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(303u); +x_1 = lean_unsigned_to_nat(304u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7813,7 +7848,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandMutualPreamble_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(316u); +x_1 = lean_unsigned_to_nat(317u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7825,7 +7860,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandMutualPreamble_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(323u); +x_1 = lean_unsigned_to_nat(324u); x_2 = lean_unsigned_to_nat(74u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7853,7 +7888,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandMutualPreamble_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(316u); +x_1 = lean_unsigned_to_nat(317u); x_2 = lean_unsigned_to_nat(4u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7865,7 +7900,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandMutualPreamble_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(316u); +x_1 = lean_unsigned_to_nat(317u); x_2 = lean_unsigned_to_nat(24u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -8021,7 +8056,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabMutual_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(326u); +x_1 = lean_unsigned_to_nat(327u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -8033,7 +8068,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabMutual_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(332u); +x_1 = lean_unsigned_to_nat(333u); x_2 = lean_unsigned_to_nat(152u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -8061,7 +8096,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabMutual_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(326u); +x_1 = lean_unsigned_to_nat(327u); x_2 = lean_unsigned_to_nat(4u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -8073,7 +8108,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabMutual_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(326u); +x_1 = lean_unsigned_to_nat(327u); x_2 = lean_unsigned_to_nat(14u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9357,7 +9392,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabAttr_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(335u); +x_1 = lean_unsigned_to_nat(336u); x_2 = lean_unsigned_to_nat(36u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9369,7 +9404,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabAttr_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(367u); +x_1 = lean_unsigned_to_nat(368u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9397,7 +9432,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabAttr_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(335u); +x_1 = lean_unsigned_to_nat(336u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9409,7 +9444,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabAttr_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(335u); +x_1 = lean_unsigned_to_nat(336u); x_2 = lean_unsigned_to_nat(48u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10047,7 +10082,7 @@ lean_inc(x_37); x_55 = l_Lean_Syntax_node3(x_37, x_54, x_41, x_51, x_53); lean_inc(x_37); x_56 = l_Lean_Syntax_node1(x_37, x_42, x_55); -x_57 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__8; +x_57 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__10; lean_inc(x_37); x_58 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_58, 0, x_37); @@ -10365,7 +10400,7 @@ x_22 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_22, 0, x_17); lean_ctor_set(x_22, 1, x_20); lean_ctor_set(x_22, 2, x_21); -x_23 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__8; +x_23 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__10; lean_inc(x_17); x_24 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_24, 0, x_17); @@ -10487,7 +10522,7 @@ lean_inc(x_17); x_79 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_79, 0, x_17); lean_ctor_set(x_79, 1, x_78); -x_80 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__12; +x_80 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__14; lean_inc(x_17); x_81 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_81, 0, x_17); @@ -10507,7 +10542,7 @@ x_89 = l_Lean_Elab_Command_expandInitialize___lambda__2___closed__9; lean_inc(x_22); lean_inc(x_17); x_90 = l_Lean_Syntax_node2(x_17, x_89, x_22, x_88); -x_91 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__13; +x_91 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__15; lean_inc(x_22); lean_inc(x_17); x_92 = l_Lean_Syntax_node4(x_17, x_91, x_81, x_87, x_90, x_22); @@ -11562,7 +11597,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandInitialize_decl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(369u); +x_1 = lean_unsigned_to_nat(370u); x_2 = lean_unsigned_to_nat(50u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11574,7 +11609,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandInitialize_decl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(381u); +x_1 = lean_unsigned_to_nat(382u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11602,7 +11637,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandInitialize_decl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(369u); +x_1 = lean_unsigned_to_nat(370u); x_2 = lean_unsigned_to_nat(54u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11614,7 +11649,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandInitialize_decl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(369u); +x_1 = lean_unsigned_to_nat(370u); x_2 = lean_unsigned_to_nat(70u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11660,7 +11695,7 @@ x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__1() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -11670,37 +11705,37 @@ 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_Declaration___hyg_7880____closed__2() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____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_Declaration___hyg_7880____closed__1; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__1; x_2 = l_Lean_Elab_Command_elabAxiom___lambda__5___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_Declaration___hyg_7880____closed__3() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____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_Declaration___hyg_7880____closed__2; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__2; x_2 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__3; 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_Declaration___hyg_7880____closed__4() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____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_Declaration___hyg_7880____closed__3; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__3; x_2 = l_Lean_Elab_Command_expandInitialize___lambda__1___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_Declaration___hyg_7880____closed__5() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__5() { _start: { lean_object* x_1; @@ -11708,37 +11743,37 @@ x_1 = lean_mk_string_from_bytes("_@", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__6() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____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_Declaration___hyg_7880____closed__4; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__5; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__4; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____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_Declaration___hyg_7880____closed__7() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____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_Declaration___hyg_7880____closed__6; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__6; x_2 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___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_Declaration___hyg_7880____closed__8() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____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_Declaration___hyg_7880____closed__7; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__7; x_2 = l_Lean_Elab_Command_elabAxiom___lambda__5___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_Declaration___hyg_7880____closed__9() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__9() { _start: { lean_object* x_1; @@ -11746,17 +11781,17 @@ x_1 = lean_mk_string_from_bytes("Declaration", 11); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__10() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____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_Declaration___hyg_7880____closed__8; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__9; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__8; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____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_Declaration___hyg_7880____closed__11() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__11() { _start: { lean_object* x_1; @@ -11764,33 +11799,33 @@ x_1 = lean_mk_string_from_bytes("_hyg", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__12() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____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_Declaration___hyg_7880____closed__10; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__11; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__10; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__11; 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_Declaration___hyg_7880____closed__13() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____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_Declaration___hyg_7880____closed__12; -x_2 = lean_unsigned_to_nat(7880u); +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__12; +x_2 = lean_unsigned_to_nat(7888u); 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_Declaration___hyg_7880_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888_(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_elabAxiom___lambda__5___closed__6; x_3 = 0; -x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__13; +x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__13; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -11892,6 +11927,10 @@ l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__20 = lean_mark_persistent(l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__20); l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__21 = _init_l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__21(); lean_mark_persistent(l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__21); +l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__22 = _init_l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__22(); +lean_mark_persistent(l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__22); +l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__23 = _init_l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__23(); +lean_mark_persistent(l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__23); l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isInstanceDef___closed__1 = _init_l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isInstanceDef___closed__1(); lean_mark_persistent(l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isInstanceDef___closed__1); l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isInstanceDef___closed__2 = _init_l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isInstanceDef___closed__2(); @@ -12395,33 +12434,33 @@ lean_mark_persistent(l___regBuiltin_Lean_Elab_Command_expandInitialize_declRange if (builtin) {res = l___regBuiltin_Lean_Elab_Command_expandInitialize_declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__1(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__1); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__2(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__2); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__3(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__3); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__4 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__4(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__4); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__5 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__5(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__5); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__6 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__6(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__6); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__7 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__7(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__7); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__8 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__8(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__8); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__9 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__9(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__9); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__10 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__10(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__10); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__11 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__11(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__11); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__12 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__12(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__12); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__13 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__13(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__13); -if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880_(lean_io_mk_world()); +}l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__1(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__1); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__2(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__2); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__3(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__3); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__4 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__4(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__4); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__5 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__5(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__5); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__6 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__6(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__6); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__7 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__7(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__7); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__8 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__8(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__8); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__9 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__9(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__9); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__10 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__10(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__10); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__11 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__11(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__11); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__12 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__12(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__12); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__13 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__13(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__13); +if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888_(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/DefView.c b/stage0/stdlib/Lean/Elab/DefView.c index 339f9f919cdb..b109160054b9 100644 --- a/stage0/stdlib/Lean/Elab/DefView.c +++ b/stage0/stdlib/Lean/Elab/DefView.c @@ -14,9 +14,9 @@ extern "C" { #endif LEAN_EXPORT lean_object* l_Lean_Meta_visitLambda_visit___at_Lean_Elab_Command_mkInstanceName___spec__5___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_mkInstanceName___spec__14___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withoutErrToSorry___at_Lean_Elab_Command_mkInstanceName___spec__13(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_DefView___hyg_2224____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_mkInstanceName___spec__14___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_mkInstanceName___spec__14___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint32_t lean_string_utf8_get(lean_object*, lean_object*); @@ -31,6 +31,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkDefViewOfExample(lean_object*, le LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkDefViewOfInstance(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forEachExpr_x27_visit___at_Lean_Elab_Command_mkInstanceName___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* l_Lean_Elab_Command_runTermElabM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__9; LEAN_EXPORT lean_object* l_Lean_Meta_forEachExpr_x27_visit___at_Lean_Elab_Command_mkInstanceName___spec__3___lambda__1(lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_maxRecDepthErrorMessage; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_mkInstanceName___spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -43,6 +44,7 @@ static lean_object* l_Lean_Elab_Command_mkInstanceName___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkDefViewOfOpaque___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkInstanceName___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_Meta_forEachExpr_x27___at_Lean_Elab_Command_mkInstanceName___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_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__11; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_mkDefViewOfInstance___spec__2(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_Elab_liftMacroM___at_Lean_Elab_Command_mkInstanceName___spec__14(lean_object*, lean_object*, lean_object*, lean_object*); @@ -53,12 +55,12 @@ static lean_object* l_Lean_Elab_Command_isDefLike___closed__5; uint8_t lean_usize_dec_eq(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Meta_visitForall_visit___at_Lean_Elab_Command_mkInstanceName___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*); static lean_object* l_Lean_Elab_Command_mkDefViewOfInstance___closed__6; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__11; static lean_object* l_Lean_Elab_Command_mkDefViewOfInstance___closed__10; lean_object* l_Lean_Syntax_getArgs(lean_object*); lean_object* l_ST_Prim_Ref_get___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkDefViewOfInstance___lambda__1___boxed(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_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_mkDefViewOfInstance___spec__5___boxed(lean_object*, lean_object*); 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_Meta_withLocalDecl___at_Lean_Elab_Command_mkInstanceName___spec__6___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*); @@ -75,11 +77,11 @@ lean_object* l_Lean_stringToMessageData(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_mkDefViewOfInstance___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_isDefLike___closed__9; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_mkInstanceName___spec__14___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__10; uint8_t lean_string_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_mkDefViewOfOpaque___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkDefViewOfOpaque___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_mkDefViewOfOpaque___lambda__2___closed__2; +static lean_object* l_Lean_Elab_Command_isDefLike___closed__12; LEAN_EXPORT uint8_t l_Lean_Elab_DefKind_isExample(uint8_t); LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkDefViewOfOpaque___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_visitLet_visit___at_Lean_Elab_Command_mkInstanceName___spec__11___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*); @@ -101,6 +103,7 @@ static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_mkInstanceN LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_mkDefViewOfInstance___spec__5(lean_object*, lean_object*); uint8_t l_Lean_Expr_isSort(lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__7; 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_Lean_Elab_liftMacroM___at_Lean_Elab_Command_mkInstanceName___spec__14___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_liftExcept___at_Lean_Elab_liftMacroM___spec__1(lean_object*, lean_object*, lean_object*); @@ -114,14 +117,13 @@ lean_object* l_Lean_ResolveName_resolveNamespace(lean_object*, lean_object*, lea LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_DefView_isInstance___spec__1(lean_object*, size_t, size_t); static lean_object* l_Lean_Elab_Command_mkDefViewOfAbbrev___closed__1; lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__15; LEAN_EXPORT uint8_t l_Lean_Elab_DefKind_isTheorem(uint8_t); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_mkInstanceName___spec__14___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_instInhabitedDefView___closed__2; lean_object* l_Lean_Elab_Term_withoutErrToSorryImp___rarg(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_DefView___hyg_2224____closed__8; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_mkDefView___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_mkInstanceName___spec__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__5; lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_mkDefViewOfInstance___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_mkInstanceName___spec__16(lean_object*, lean_object*, lean_object*, lean_object*); @@ -135,9 +137,9 @@ LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_mkInstanc lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_visitLambda_visit___at_Lean_Elab_Command_mkInstanceName___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_mkInstanceName___spec__15(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__13; LEAN_EXPORT lean_object* l_Lean_Elab_DefKind_toCtorIdx___boxed(lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__16; LEAN_EXPORT uint8_t l_Lean_Elab_Command_isDefLike(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_mkDefViewOfInstance___spec__8(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getOptional_x3f(lean_object*); @@ -151,7 +153,6 @@ lean_object* l_Lean_Expr_constName_x21(lean_object*); lean_object* l_Lean_Elab_Command_getCurrMacroScope(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_HashMapImp_find_x3f___at_Lean_ForEachExpr_visit___spec__1(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Elab_DefView_isInstance(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224_(lean_object*); static lean_object* l_Lean_Elab_Command_mkDefViewOfExample___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_DefKind_toCtorIdx(uint8_t); static lean_object* l_Lean_Elab_Command_mkDefViewOfAbbrev___closed__4; @@ -168,25 +169,19 @@ LEAN_EXPORT lean_object* l_Lean_Elab_DefKind_isExample___boxed(lean_object*); 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*); LEAN_EXPORT lean_object* l_Lean_Meta_visitLet___at_Lean_Elab_Command_mkInstanceName___spec__10(lean_object*, lean_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_isDefLike___closed__13; LEAN_EXPORT lean_object* l_Lean_Elab_instInhabitedDefView; LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkInstanceName___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__9; lean_object* l_Lean_Elab_Command_getScope___rarg(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__3; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__4; static lean_object* l_Lean_Meta_forEachExpr_x27___at_Lean_Elab_Command_mkInstanceName___spec__2___closed__1; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__7; static lean_object* l_Lean_Elab_Command_mkDefViewOfInstance___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_forEachExpr___at_Lean_Elab_Command_mkInstanceName___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__12; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__6; uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_visitForall_visit___at_Lean_Elab_Command_mkInstanceName___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* l_Lean_Elab_expandOptNamedPrio___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__17; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_visitLet_visit___at_Lean_Elab_Command_mkInstanceName___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* 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_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__14; LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkDefView(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkInstanceName___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_isDefLike___boxed(lean_object*); @@ -210,6 +205,7 @@ uint8_t lean_nat_dec_lt(lean_object*, lean_object*); lean_object* lean_environment_main_module(lean_object*); static lean_object* l_Lean_Elab_instInhabitedDefView___closed__3; static lean_object* l_Lean_Elab_Command_mkDefViewOfOpaque___lambda__2___closed__4; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__10; 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_mkDefViewOfInstance___closed__8; @@ -217,10 +213,10 @@ 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*); static lean_object* l_Lean_Elab_Command_isDefLike___closed__8; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__15; lean_object* l_Lean_Syntax_node1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_mkDefViewOfExample___closed__2; LEAN_EXPORT uint8_t l___private_Lean_Elab_DefView_0__Lean_Elab_beqDefKind____x40_Lean_Elab_DefView___hyg_15_(uint8_t, uint8_t); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__2; uint8_t l_Lean_Syntax_isNone(lean_object*); static lean_object* l_Lean_Elab_Command_mkDefViewOfOpaque___lambda__2___closed__5; static lean_object* l_Lean_Elab_Command_isDefLike___closed__6; @@ -231,9 +227,10 @@ static lean_object* l_Lean_Elab_Command_mkDefViewOfInstance___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_mkDefViewOfInstance___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_mkInstanceName___lambda__2___closed__1; uint8_t l_Lean_Expr_isProp(lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__12; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_Command_mkInstanceName___spec__6(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__6; lean_object* l_Lean_Elab_expandOptDeclSig(lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__13; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_mkDefViewOfInstance___spec__10___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Elab_Command_mkInstanceName___spec__12___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_erase_macro_scopes(lean_object*); @@ -250,6 +247,8 @@ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_DefView_isInstanc lean_object* l_String_capitalize(lean_object*); static lean_object* l_Lean_Elab_Command_isDefLike___closed__4; size_t lean_usize_add(size_t, size_t); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__3; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__4; static lean_object* l_Lean_Elab_Command_mkDefViewOfOpaque___lambda__2___closed__6; static lean_object* l_Lean_Elab_Command_mkDefViewOfOpaque___closed__2; lean_object* lean_array_uget(lean_object*, size_t); @@ -258,11 +257,12 @@ lean_object* l_Lean_Elab_mkUnusedBaseName(lean_object*, lean_object*, lean_objec 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*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__14; lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_mkInstanceName___lambda__1___closed__2; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__1; static lean_object* l_Lean_Elab_Command_mkDefViewOfInstance___closed__11; LEAN_EXPORT lean_object* l_Lean_Elab_DefKind_noConfusion___rarg___lambda__1(lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__16; LEAN_EXPORT lean_object* l___private_Lean_Elab_DefView_0__Lean_Elab_beqDefKind____x40_Lean_Elab_DefView___hyg_15____boxed(lean_object*, lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_DefKind_isTheorem___boxed(lean_object*); @@ -295,6 +295,7 @@ lean_object* l_Lean_Elab_Term_elabType(lean_object*, lean_object*, lean_object*, lean_object* l_Lean_Syntax_mkNumLit(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_Meta_withLocalDecl___at_Lean_Elab_Command_mkInstanceName___spec__6___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*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkDefViewOfDef(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_DefKind_toCtorIdx(uint8_t x_1) { _start: @@ -6086,7 +6087,7 @@ static lean_object* _init_l_Lean_Elab_Command_isDefLike___closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("def", 3); +x_1 = lean_mk_string_from_bytes("definition", 10); return x_1; } } @@ -6106,7 +6107,7 @@ static lean_object* _init_l_Lean_Elab_Command_isDefLike___closed__5() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("theorem", 7); +x_1 = lean_mk_string_from_bytes("def", 3); return x_1; } } @@ -6126,7 +6127,7 @@ static lean_object* _init_l_Lean_Elab_Command_isDefLike___closed__7() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("opaque", 6); +x_1 = lean_mk_string_from_bytes("theorem", 7); return x_1; } } @@ -6145,6 +6146,26 @@ return x_5; static lean_object* _init_l_Lean_Elab_Command_isDefLike___closed__9() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("opaque", 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Command_isDefLike___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_Elab_Command_mkDefViewOfInstance___closed__1; +x_2 = l_Lean_Elab_Command_mkDefViewOfInstance___closed__2; +x_3 = l_Lean_Elab_Command_mkDefViewOfInstance___closed__7; +x_4 = l_Lean_Elab_Command_isDefLike___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_Elab_Command_isDefLike___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_Elab_Command_mkDefViewOfInstance___closed__1; x_2 = l_Lean_Elab_Command_mkDefViewOfInstance___closed__2; @@ -6154,7 +6175,7 @@ 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_isDefLike___closed__10() { +static lean_object* _init_l_Lean_Elab_Command_isDefLike___closed__12() { _start: { lean_object* x_1; @@ -6162,14 +6183,14 @@ x_1 = lean_mk_string_from_bytes("example", 7); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_isDefLike___closed__11() { +static lean_object* _init_l_Lean_Elab_Command_isDefLike___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_Elab_Command_mkDefViewOfInstance___closed__1; x_2 = l_Lean_Elab_Command_mkDefViewOfInstance___closed__2; x_3 = l_Lean_Elab_Command_mkDefViewOfInstance___closed__7; -x_4 = l_Lean_Elab_Command_isDefLike___closed__10; +x_4 = l_Lean_Elab_Command_isDefLike___closed__12; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } @@ -6199,32 +6220,21 @@ x_10 = lean_name_eq(x_2, x_9); if (x_10 == 0) { lean_object* x_11; uint8_t x_12; -x_11 = l_Lean_Elab_Command_isDefLike___closed__9; +x_11 = l_Lean_Elab_Command_isDefLike___closed__10; x_12 = lean_name_eq(x_2, x_11); if (x_12 == 0) { lean_object* x_13; uint8_t x_14; x_13 = l_Lean_Elab_Command_isDefLike___closed__11; x_14 = lean_name_eq(x_2, x_13); -lean_dec(x_2); -return x_14; -} -else -{ -uint8_t x_15; -lean_dec(x_2); -x_15 = 1; -return x_15; -} -} -else +if (x_14 == 0) { -uint8_t x_16; +lean_object* x_15; uint8_t x_16; +x_15 = l_Lean_Elab_Command_isDefLike___closed__13; +x_16 = lean_name_eq(x_2, x_15); lean_dec(x_2); -x_16 = 1; return x_16; } -} else { uint8_t x_17; @@ -6249,6 +6259,30 @@ x_19 = 1; return x_19; } } +else +{ +uint8_t x_20; +lean_dec(x_2); +x_20 = 1; +return x_20; +} +} +else +{ +uint8_t x_21; +lean_dec(x_2); +x_21 = 1; +return x_21; +} +} +else +{ +uint8_t x_22; +lean_dec(x_2); +x_22 = 1; +return x_22; +} +} } LEAN_EXPORT lean_object* l_Lean_Elab_Command_isDefLike___boxed(lean_object* x_1) { _start: @@ -6355,62 +6389,54 @@ x_14 = lean_name_eq(x_6, x_13); if (x_14 == 0) { lean_object* x_15; uint8_t x_16; -x_15 = l_Lean_Elab_Command_isDefLike___closed__9; +x_15 = l_Lean_Elab_Command_isDefLike___closed__10; x_16 = lean_name_eq(x_6, x_15); if (x_16 == 0) { lean_object* x_17; uint8_t x_18; x_17 = l_Lean_Elab_Command_isDefLike___closed__11; x_18 = lean_name_eq(x_6, x_17); -lean_dec(x_6); if (x_18 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_19; uint8_t x_20; +x_19 = l_Lean_Elab_Command_isDefLike___closed__13; +x_20 = lean_name_eq(x_6, x_19); +lean_dec(x_6); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; lean_dec(x_2); lean_dec(x_1); -x_19 = l_Lean_Elab_Command_mkDefView___closed__2; -x_20 = l_Lean_throwError___at_Lean_Elab_Command_mkDefView___spec__1(x_19, x_3, x_4, x_5); +x_21 = l_Lean_Elab_Command_mkDefView___closed__2; +x_22 = l_Lean_throwError___at_Lean_Elab_Command_mkDefView___spec__1(x_21, x_3, x_4, x_5); lean_dec(x_4); -return x_20; +return x_22; } else { -lean_object* x_21; lean_object* x_22; +lean_object* x_23; lean_object* x_24; lean_dec(x_4); lean_dec(x_3); -x_21 = l_Lean_Elab_Command_mkDefViewOfExample(x_1, x_2); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_5); -return x_22; -} -} -else -{ -lean_object* x_23; -lean_dec(x_6); -x_23 = l_Lean_Elab_Command_mkDefViewOfInstance(x_1, x_2, x_3, x_4, x_5); -return x_23; +x_23 = l_Lean_Elab_Command_mkDefViewOfExample(x_1, x_2); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_5); +return x_24; } } else { -lean_object* x_24; +lean_object* x_25; lean_dec(x_6); -x_24 = l_Lean_Elab_Command_mkDefViewOfOpaque(x_1, x_2, x_3, x_4, x_5); -return x_24; +x_25 = l_Lean_Elab_Command_mkDefViewOfInstance(x_1, x_2, x_3, x_4, x_5); +return x_25; } } else { -lean_object* x_25; lean_object* x_26; +lean_object* x_26; lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_3); -x_25 = l_Lean_Elab_Command_mkDefViewOfTheorem(x_1, x_2); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_5); +x_26 = l_Lean_Elab_Command_mkDefViewOfOpaque(x_1, x_2, x_3, x_4, x_5); return x_26; } } @@ -6420,7 +6446,7 @@ lean_object* x_27; lean_object* x_28; lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); -x_27 = l_Lean_Elab_Command_mkDefViewOfDef(x_1, x_2); +x_27 = l_Lean_Elab_Command_mkDefViewOfTheorem(x_1, x_2); x_28 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_28, 0, x_27); lean_ctor_set(x_28, 1, x_5); @@ -6433,13 +6459,39 @@ lean_object* x_29; lean_object* x_30; lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); -x_29 = l_Lean_Elab_Command_mkDefViewOfAbbrev(x_1, x_2); +x_29 = l_Lean_Elab_Command_mkDefViewOfDef(x_1, x_2); x_30 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_30, 0, x_29); lean_ctor_set(x_30, 1, x_5); return x_30; } } +else +{ +lean_object* x_31; lean_object* x_32; +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +x_31 = l_Lean_Elab_Command_mkDefViewOfDef(x_1, x_2); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_5); +return x_32; +} +} +else +{ +lean_object* x_33; lean_object* x_34; +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +x_33 = l_Lean_Elab_Command_mkDefViewOfAbbrev(x_1, x_2); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_5); +return x_34; +} +} } LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_mkDefView___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: @@ -6450,7 +6502,7 @@ lean_dec(x_3); return x_5; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__1() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__1() { _start: { lean_object* x_1; @@ -6458,25 +6510,17 @@ x_1 = lean_mk_string_from_bytes("Elab", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("definition", 10); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__3() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____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_DefView___hyg_2224____closed__1; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__2; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__1; +x_2 = l_Lean_Elab_Command_isDefLike___closed__3; 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_DefView___hyg_2224____closed__4() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -6486,27 +6530,27 @@ 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_DefView___hyg_2224____closed__5() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__4; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__1; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__3; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____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_DefView___hyg_2224____closed__6() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____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_DefView___hyg_2224____closed__5; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__4; x_2 = l_Lean_Elab_Command_mkDefViewOfInstance___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_DefView___hyg_2224____closed__7() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__6() { _start: { lean_object* x_1; @@ -6514,17 +6558,17 @@ x_1 = lean_mk_string_from_bytes("initFn", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__8() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____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_DefView___hyg_2224____closed__6; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__7; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__5; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____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_DefView___hyg_2224____closed__9() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__8() { _start: { lean_object* x_1; @@ -6532,37 +6576,37 @@ x_1 = lean_mk_string_from_bytes("_@", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__10() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____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_DefView___hyg_2224____closed__8; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__9; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__7; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____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_DefView___hyg_2224____closed__11() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____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_DefView___hyg_2224____closed__10; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__9; x_2 = l_Lean_Elab_Command_mkDefViewOfInstance___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_DefView___hyg_2224____closed__12() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____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_DefView___hyg_2224____closed__11; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__1; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__10; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____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_DefView___hyg_2224____closed__13() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__12() { _start: { lean_object* x_1; @@ -6570,17 +6614,17 @@ x_1 = lean_mk_string_from_bytes("DefView", 7); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__14() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____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_DefView___hyg_2224____closed__12; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__13; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__11; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____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_DefView___hyg_2224____closed__15() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__14() { _start: { lean_object* x_1; @@ -6588,33 +6632,33 @@ x_1 = lean_mk_string_from_bytes("_hyg", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__16() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__14; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__15; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__13; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__14; 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_DefView___hyg_2224____closed__17() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____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_DefView___hyg_2224____closed__16; -x_2 = lean_unsigned_to_nat(2224u); +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__15; +x_2 = lean_unsigned_to_nat(2255u); 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_DefView___hyg_2224_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255_(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_DefView___hyg_2224____closed__3; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__2; x_3 = 0; -x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__17; +x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__16; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -6765,45 +6809,47 @@ l_Lean_Elab_Command_isDefLike___closed__10 = _init_l_Lean_Elab_Command_isDefLike lean_mark_persistent(l_Lean_Elab_Command_isDefLike___closed__10); l_Lean_Elab_Command_isDefLike___closed__11 = _init_l_Lean_Elab_Command_isDefLike___closed__11(); lean_mark_persistent(l_Lean_Elab_Command_isDefLike___closed__11); +l_Lean_Elab_Command_isDefLike___closed__12 = _init_l_Lean_Elab_Command_isDefLike___closed__12(); +lean_mark_persistent(l_Lean_Elab_Command_isDefLike___closed__12); +l_Lean_Elab_Command_isDefLike___closed__13 = _init_l_Lean_Elab_Command_isDefLike___closed__13(); +lean_mark_persistent(l_Lean_Elab_Command_isDefLike___closed__13); l_Lean_Elab_Command_mkDefView___closed__1 = _init_l_Lean_Elab_Command_mkDefView___closed__1(); lean_mark_persistent(l_Lean_Elab_Command_mkDefView___closed__1); l_Lean_Elab_Command_mkDefView___closed__2 = _init_l_Lean_Elab_Command_mkDefView___closed__2(); lean_mark_persistent(l_Lean_Elab_Command_mkDefView___closed__2); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__1(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__1); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__2(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__2); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__3(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__3); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__4 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__4(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__4); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__5 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__5(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__5); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__6 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__6(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__6); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__7 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__7(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__7); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__8 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__8(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__8); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__9 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__9(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__9); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__10 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__10(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__10); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__11 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__11(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__11); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__12 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__12(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__12); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__13 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__13(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__13); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__14 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__14(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__14); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__15 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__15(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__15); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__16 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__16(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__16); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__17 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__17(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__17); -if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224_(lean_io_mk_world()); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__1(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__1); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__2(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__2); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__3(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__3); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__4 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__4(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__4); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__5 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__5(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__5); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__6 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__6(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__6); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__7 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__7(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__7); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__8 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__8(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__8); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__9 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__9(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__9); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__10 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__10(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__10); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__11 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__11(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__11); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__12 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__12(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__12); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__13 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__13(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__13); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__14 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__14(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__14); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__15 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__15(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__15); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__16 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__16(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__16); +if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Elab/Deriving/BEq.c b/stage0/stdlib/Lean/Elab/Deriving/BEq.c index a4dadc7867b5..dbf7f76c2208 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/BEq.c +++ b/stage0/stdlib/Lean/Elab/Deriving/BEq.c @@ -27,6 +27,7 @@ static lean_object* l_Lean_Elab_Deriving_BEq_initFn____x40_Lean_Elab_Deriving_BE static lean_object* l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__22; static lean_object* l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__2; lean_object* lean_mk_empty_array_with_capacity(lean_object*); +static lean_object* l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__31; static lean_object* l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__9; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__15; static lean_object* l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__13; @@ -2884,7 +2885,7 @@ static lean_object* _init_l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___c _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("def", 3); +x_1 = lean_mk_string_from_bytes("definition", 10); return x_1; } } @@ -2904,23 +2905,31 @@ static lean_object* _init_l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___c _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("declId", 6); +x_1 = lean_mk_string_from_bytes("def", 3); return x_1; } } static lean_object* _init_l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__11() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("declId", 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_BEq_mkAuxFunction___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_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__1; x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__2; x_3 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__1; -x_4 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__10; +x_4 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___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_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__12() { +static lean_object* _init_l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__13() { _start: { lean_object* x_1; @@ -2928,19 +2937,19 @@ x_1 = lean_mk_string_from_bytes("optDeclSig", 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__13() { +static lean_object* _init_l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___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_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__1; x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__2; x_3 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__1; -x_4 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__12; +x_4 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___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_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__14() { +static lean_object* _init_l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__15() { _start: { lean_object* x_1; @@ -2948,19 +2957,19 @@ x_1 = lean_mk_string_from_bytes("typeSpec", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__15() { +static lean_object* _init_l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___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_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__1; x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__2; x_3 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__3; -x_4 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__14; +x_4 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___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_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__16() { +static lean_object* _init_l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__17() { _start: { lean_object* x_1; @@ -2968,7 +2977,7 @@ x_1 = lean_mk_string_from_bytes(":", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__17() { +static lean_object* _init_l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__18() { _start: { lean_object* x_1; lean_object* x_2; @@ -2977,7 +2986,7 @@ x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__18() { +static lean_object* _init_l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -2987,53 +2996,53 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__19() { +static lean_object* _init_l_Lean_Elab_Deriving_BEq_mkAuxFunction___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_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__18; +x_2 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___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_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__20() { +static lean_object* _init_l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__21() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__18; +x_1 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___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_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__21() { +static lean_object* _init_l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__20; +x_2 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__21; 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_BEq_mkAuxFunction___lambda__1___closed__22() { +static lean_object* _init_l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__19; -x_2 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__21; +x_1 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__20; +x_2 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___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_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__23() { +static lean_object* _init_l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__24() { _start: { lean_object* x_1; @@ -3041,19 +3050,19 @@ x_1 = lean_mk_string_from_bytes("declValSimple", 13); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__24() { +static lean_object* _init_l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__25() { _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_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__1; x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__2; x_3 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__1; -x_4 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__23; +x_4 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__24; 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_BEq_mkAuxFunction___lambda__1___closed__25() { +static lean_object* _init_l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__26() { _start: { lean_object* x_1; @@ -3061,7 +3070,7 @@ x_1 = lean_mk_string_from_bytes(":=", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__26() { +static lean_object* _init_l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__27() { _start: { lean_object* x_1; @@ -3069,7 +3078,7 @@ x_1 = lean_mk_string_from_bytes("Termination", 11); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__27() { +static lean_object* _init_l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__28() { _start: { lean_object* x_1; @@ -3077,19 +3086,19 @@ x_1 = lean_mk_string_from_bytes("suffix", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__28() { +static lean_object* _init_l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__29() { _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_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__1; x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__2; -x_3 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__26; -x_4 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__27; +x_3 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__27; +x_4 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__28; 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_BEq_mkAuxFunction___lambda__1___closed__29() { +static lean_object* _init_l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__30() { _start: { lean_object* x_1; @@ -3097,14 +3106,14 @@ x_1 = lean_mk_string_from_bytes("partial", 7); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__30() { +static lean_object* _init_l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__31() { _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_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__1; x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__2; x_3 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__1; -x_4 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__29; +x_4 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__30; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } @@ -3158,13 +3167,13 @@ x_32 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__5; lean_inc_n(x_26, 5); lean_inc(x_17); x_33 = l_Lean_Syntax_node6(x_17, x_32, x_26, x_26, x_31, x_26, x_26, x_26); -x_34 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__8; +x_34 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__10; lean_inc(x_17); x_35 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_35, 0, x_17); lean_ctor_set(x_35, 1, x_34); x_36 = lean_mk_syntax_ident(x_3); -x_37 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__11; +x_37 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__12; lean_inc(x_26); lean_inc(x_17); x_38 = l_Lean_Syntax_node2(x_17, x_37, x_36, x_26); @@ -3174,39 +3183,39 @@ x_40 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_40, 0, x_17); lean_ctor_set(x_40, 1, x_24); lean_ctor_set(x_40, 2, x_39); -x_41 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__16; +x_41 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__17; lean_inc(x_17); x_42 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_42, 0, x_17); lean_ctor_set(x_42, 1, x_41); -x_43 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__18; +x_43 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__19; x_44 = l_Lean_addMacroScope(x_23, x_43, x_18); -x_45 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__17; -x_46 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__22; +x_45 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__18; +x_46 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__23; lean_inc(x_17); x_47 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_47, 0, x_17); lean_ctor_set(x_47, 1, x_45); lean_ctor_set(x_47, 2, x_44); lean_ctor_set(x_47, 3, x_46); -x_48 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__15; +x_48 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__16; lean_inc(x_17); x_49 = l_Lean_Syntax_node2(x_17, x_48, x_42, x_47); lean_inc(x_17); x_50 = l_Lean_Syntax_node1(x_17, x_24, x_49); -x_51 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__13; +x_51 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__14; lean_inc(x_17); x_52 = l_Lean_Syntax_node2(x_17, x_51, x_40, x_50); -x_53 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__25; +x_53 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__26; lean_inc(x_17); x_54 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_54, 0, x_17); lean_ctor_set(x_54, 1, x_53); -x_55 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__28; +x_55 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__29; lean_inc_n(x_26, 2); lean_inc(x_17); x_56 = l_Lean_Syntax_node2(x_17, x_55, x_26, x_26); -x_57 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__24; +x_57 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__25; lean_inc(x_26); lean_inc(x_17); x_58 = l_Lean_Syntax_node4(x_17, x_57, x_54, x_4, x_56, x_26); @@ -3251,13 +3260,13 @@ x_75 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__5; lean_inc_n(x_69, 5); lean_inc(x_17); x_76 = l_Lean_Syntax_node6(x_17, x_75, x_69, x_69, x_74, x_69, x_69, x_69); -x_77 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__8; +x_77 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__10; lean_inc(x_17); x_78 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_78, 0, x_17); lean_ctor_set(x_78, 1, x_77); x_79 = lean_mk_syntax_ident(x_3); -x_80 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__11; +x_80 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__12; lean_inc(x_69); lean_inc(x_17); x_81 = l_Lean_Syntax_node2(x_17, x_80, x_79, x_69); @@ -3267,39 +3276,39 @@ x_83 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_83, 0, x_17); lean_ctor_set(x_83, 1, x_67); lean_ctor_set(x_83, 2, x_82); -x_84 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__16; +x_84 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__17; lean_inc(x_17); x_85 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_85, 0, x_17); lean_ctor_set(x_85, 1, x_84); -x_86 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__18; +x_86 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__19; x_87 = l_Lean_addMacroScope(x_66, x_86, x_18); -x_88 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__17; -x_89 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__22; +x_88 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__18; +x_89 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__23; lean_inc(x_17); x_90 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_90, 0, x_17); lean_ctor_set(x_90, 1, x_88); lean_ctor_set(x_90, 2, x_87); lean_ctor_set(x_90, 3, x_89); -x_91 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__15; +x_91 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__16; lean_inc(x_17); x_92 = l_Lean_Syntax_node2(x_17, x_91, x_85, x_90); lean_inc(x_17); x_93 = l_Lean_Syntax_node1(x_17, x_67, x_92); -x_94 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__13; +x_94 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__14; lean_inc(x_17); x_95 = l_Lean_Syntax_node2(x_17, x_94, x_83, x_93); -x_96 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__25; +x_96 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__26; lean_inc(x_17); x_97 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_97, 0, x_17); lean_ctor_set(x_97, 1, x_96); -x_98 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__28; +x_98 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__29; lean_inc_n(x_69, 2); lean_inc(x_17); x_99 = l_Lean_Syntax_node2(x_17, x_98, x_69, x_69); -x_100 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__24; +x_100 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__25; lean_inc(x_69); lean_inc(x_17); x_101 = l_Lean_Syntax_node4(x_17, x_100, x_97, x_4, x_99, x_69); @@ -3354,12 +3363,12 @@ lean_inc(x_110); x_123 = l_Lean_Syntax_node1(x_110, x_122, x_121); lean_inc(x_110); x_124 = l_Lean_Syntax_node1(x_110, x_117, x_123); -x_125 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__29; +x_125 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__30; lean_inc(x_110); x_126 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_126, 0, x_110); lean_ctor_set(x_126, 1, x_125); -x_127 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__30; +x_127 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__31; lean_inc(x_110); x_128 = l_Lean_Syntax_node1(x_110, x_127, x_126); lean_inc(x_110); @@ -3368,13 +3377,13 @@ x_130 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__5; lean_inc_n(x_119, 4); lean_inc(x_110); x_131 = l_Lean_Syntax_node6(x_110, x_130, x_119, x_119, x_124, x_119, x_119, x_129); -x_132 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__8; +x_132 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__10; lean_inc(x_110); x_133 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_133, 0, x_110); lean_ctor_set(x_133, 1, x_132); x_134 = lean_mk_syntax_ident(x_3); -x_135 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__11; +x_135 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__12; lean_inc(x_119); lean_inc(x_110); x_136 = l_Lean_Syntax_node2(x_110, x_135, x_134, x_119); @@ -3384,39 +3393,39 @@ x_138 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_138, 0, x_110); lean_ctor_set(x_138, 1, x_117); lean_ctor_set(x_138, 2, x_137); -x_139 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__16; +x_139 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__17; lean_inc(x_110); x_140 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_140, 0, x_110); lean_ctor_set(x_140, 1, x_139); -x_141 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__18; +x_141 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__19; x_142 = l_Lean_addMacroScope(x_116, x_141, x_111); -x_143 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__17; -x_144 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__22; +x_143 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__18; +x_144 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__23; lean_inc(x_110); x_145 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_145, 0, x_110); lean_ctor_set(x_145, 1, x_143); lean_ctor_set(x_145, 2, x_142); lean_ctor_set(x_145, 3, x_144); -x_146 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__15; +x_146 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__16; lean_inc(x_110); x_147 = l_Lean_Syntax_node2(x_110, x_146, x_140, x_145); lean_inc(x_110); x_148 = l_Lean_Syntax_node1(x_110, x_117, x_147); -x_149 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__13; +x_149 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__14; lean_inc(x_110); x_150 = l_Lean_Syntax_node2(x_110, x_149, x_138, x_148); -x_151 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__25; +x_151 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__26; lean_inc(x_110); x_152 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_152, 0, x_110); lean_ctor_set(x_152, 1, x_151); -x_153 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__28; +x_153 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__29; lean_inc_n(x_119, 2); lean_inc(x_110); x_154 = l_Lean_Syntax_node2(x_110, x_153, x_119, x_119); -x_155 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__24; +x_155 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__25; lean_inc(x_119); lean_inc(x_110); x_156 = l_Lean_Syntax_node4(x_110, x_155, x_152, x_4, x_154, x_119); @@ -3457,12 +3466,12 @@ lean_inc(x_110); x_171 = l_Lean_Syntax_node1(x_110, x_170, x_169); lean_inc(x_110); x_172 = l_Lean_Syntax_node1(x_110, x_165, x_171); -x_173 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__29; +x_173 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__30; lean_inc(x_110); x_174 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_174, 0, x_110); lean_ctor_set(x_174, 1, x_173); -x_175 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__30; +x_175 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__31; lean_inc(x_110); x_176 = l_Lean_Syntax_node1(x_110, x_175, x_174); lean_inc(x_110); @@ -3471,13 +3480,13 @@ x_178 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__5; lean_inc_n(x_167, 4); lean_inc(x_110); x_179 = l_Lean_Syntax_node6(x_110, x_178, x_167, x_167, x_172, x_167, x_167, x_177); -x_180 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__8; +x_180 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__10; lean_inc(x_110); x_181 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_181, 0, x_110); lean_ctor_set(x_181, 1, x_180); x_182 = lean_mk_syntax_ident(x_3); -x_183 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__11; +x_183 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__12; lean_inc(x_167); lean_inc(x_110); x_184 = l_Lean_Syntax_node2(x_110, x_183, x_182, x_167); @@ -3487,39 +3496,39 @@ x_186 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_186, 0, x_110); lean_ctor_set(x_186, 1, x_165); lean_ctor_set(x_186, 2, x_185); -x_187 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__16; +x_187 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__17; lean_inc(x_110); x_188 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_188, 0, x_110); lean_ctor_set(x_188, 1, x_187); -x_189 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__18; +x_189 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__19; x_190 = l_Lean_addMacroScope(x_164, x_189, x_111); -x_191 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__17; -x_192 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__22; +x_191 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__18; +x_192 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__23; lean_inc(x_110); x_193 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_193, 0, x_110); lean_ctor_set(x_193, 1, x_191); lean_ctor_set(x_193, 2, x_190); lean_ctor_set(x_193, 3, x_192); -x_194 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__15; +x_194 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__16; lean_inc(x_110); x_195 = l_Lean_Syntax_node2(x_110, x_194, x_188, x_193); lean_inc(x_110); x_196 = l_Lean_Syntax_node1(x_110, x_165, x_195); -x_197 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__13; +x_197 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__14; lean_inc(x_110); x_198 = l_Lean_Syntax_node2(x_110, x_197, x_186, x_196); -x_199 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__25; +x_199 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__26; lean_inc(x_110); x_200 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_200, 0, x_110); lean_ctor_set(x_200, 1, x_199); -x_201 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__28; +x_201 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__29; lean_inc_n(x_167, 2); lean_inc(x_110); x_202 = l_Lean_Syntax_node2(x_110, x_201, x_167, x_167); -x_203 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__24; +x_203 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__25; lean_inc(x_167); lean_inc(x_110); x_204 = l_Lean_Syntax_node4(x_110, x_203, x_200, x_4, x_202, x_167); @@ -4703,7 +4712,7 @@ static lean_object* _init_l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Derivi { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__18; +x_2 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___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); @@ -4715,7 +4724,7 @@ static lean_object* _init_l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Derivi { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkBEqEnumFun___closed__11; -x_2 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__21; +x_2 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___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); @@ -4849,13 +4858,13 @@ x_32 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__5; lean_inc_n(x_26, 5); lean_inc(x_16); x_33 = l_Lean_Syntax_node6(x_16, x_32, x_26, x_26, x_31, x_26, x_26, x_26); -x_34 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__8; +x_34 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__10; lean_inc(x_16); x_35 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_35, 0, x_16); lean_ctor_set(x_35, 1, x_34); x_36 = lean_mk_syntax_ident(x_19); -x_37 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__11; +x_37 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__12; lean_inc(x_26); lean_inc(x_16); x_38 = l_Lean_Syntax_node2(x_16, x_37, x_36, x_26); @@ -4889,7 +4898,7 @@ lean_ctor_set(x_49, 2, x_47); lean_ctor_set(x_49, 3, x_43); lean_inc(x_16); x_50 = l_Lean_Syntax_node2(x_16, x_24, x_45, x_49); -x_51 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__16; +x_51 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__17; lean_inc(x_16); x_52 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_52, 0, x_16); @@ -4909,11 +4918,11 @@ lean_inc(x_16); x_58 = l_Lean_Syntax_node5(x_16, x_57, x_40, x_50, x_54, x_26, x_56); lean_inc(x_16); x_59 = l_Lean_Syntax_node1(x_16, x_24, x_58); -x_60 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__18; +x_60 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__19; lean_inc(x_17); lean_inc(x_23); x_61 = l_Lean_addMacroScope(x_23, x_60, x_17); -x_62 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__17; +x_62 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__18; x_63 = l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkBEqEnumFun___closed__12; lean_inc(x_16); x_64 = lean_alloc_ctor(3, 4, 0); @@ -4921,15 +4930,15 @@ lean_ctor_set(x_64, 0, x_16); lean_ctor_set(x_64, 1, x_62); lean_ctor_set(x_64, 2, x_61); lean_ctor_set(x_64, 3, x_63); -x_65 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__15; +x_65 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__16; lean_inc(x_16); x_66 = l_Lean_Syntax_node2(x_16, x_65, x_52, x_64); lean_inc(x_16); x_67 = l_Lean_Syntax_node1(x_16, x_24, x_66); -x_68 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__13; +x_68 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__14; lean_inc(x_16); x_69 = l_Lean_Syntax_node2(x_16, x_68, x_59, x_67); -x_70 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__25; +x_70 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__26; lean_inc(x_16); x_71 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_71, 0, x_16); @@ -4962,11 +4971,11 @@ lean_ctor_set(x_81, 3, x_43); x_82 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___lambda__1___closed__5; lean_inc(x_16); x_83 = l_Lean_Syntax_node3(x_16, x_82, x_75, x_77, x_81); -x_84 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__28; +x_84 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__29; lean_inc_n(x_26, 2); lean_inc(x_16); x_85 = l_Lean_Syntax_node2(x_16, x_84, x_26, x_26); -x_86 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__24; +x_86 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__25; lean_inc(x_26); lean_inc(x_16); x_87 = l_Lean_Syntax_node4(x_16, x_86, x_71, x_83, x_85, x_26); @@ -5011,13 +5020,13 @@ x_104 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__5; lean_inc_n(x_98, 5); lean_inc(x_16); x_105 = l_Lean_Syntax_node6(x_16, x_104, x_98, x_98, x_103, x_98, x_98, x_98); -x_106 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__8; +x_106 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__10; lean_inc(x_16); x_107 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_107, 0, x_16); lean_ctor_set(x_107, 1, x_106); x_108 = lean_mk_syntax_ident(x_19); -x_109 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__11; +x_109 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__12; lean_inc(x_98); lean_inc(x_16); x_110 = l_Lean_Syntax_node2(x_16, x_109, x_108, x_98); @@ -5051,7 +5060,7 @@ lean_ctor_set(x_121, 2, x_119); lean_ctor_set(x_121, 3, x_115); lean_inc(x_16); x_122 = l_Lean_Syntax_node2(x_16, x_96, x_117, x_121); -x_123 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__16; +x_123 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__17; lean_inc(x_16); x_124 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_124, 0, x_16); @@ -5071,11 +5080,11 @@ lean_inc(x_16); x_130 = l_Lean_Syntax_node5(x_16, x_129, x_112, x_122, x_126, x_98, x_128); lean_inc(x_16); x_131 = l_Lean_Syntax_node1(x_16, x_96, x_130); -x_132 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__18; +x_132 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__19; lean_inc(x_17); lean_inc(x_95); x_133 = l_Lean_addMacroScope(x_95, x_132, x_17); -x_134 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__17; +x_134 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__18; x_135 = l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkBEqEnumFun___closed__12; lean_inc(x_16); x_136 = lean_alloc_ctor(3, 4, 0); @@ -5083,15 +5092,15 @@ lean_ctor_set(x_136, 0, x_16); lean_ctor_set(x_136, 1, x_134); lean_ctor_set(x_136, 2, x_133); lean_ctor_set(x_136, 3, x_135); -x_137 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__15; +x_137 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__16; lean_inc(x_16); x_138 = l_Lean_Syntax_node2(x_16, x_137, x_124, x_136); lean_inc(x_16); x_139 = l_Lean_Syntax_node1(x_16, x_96, x_138); -x_140 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__13; +x_140 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__14; lean_inc(x_16); x_141 = l_Lean_Syntax_node2(x_16, x_140, x_131, x_139); -x_142 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__25; +x_142 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__26; lean_inc(x_16); x_143 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_143, 0, x_16); @@ -5124,11 +5133,11 @@ lean_ctor_set(x_153, 3, x_115); x_154 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___lambda__1___closed__5; lean_inc(x_16); x_155 = l_Lean_Syntax_node3(x_16, x_154, x_147, x_149, x_153); -x_156 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__28; +x_156 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__29; lean_inc_n(x_98, 2); lean_inc(x_16); x_157 = l_Lean_Syntax_node2(x_16, x_156, x_98, x_98); -x_158 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__24; +x_158 = l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__25; lean_inc(x_98); lean_inc(x_16); x_159 = l_Lean_Syntax_node4(x_16, x_158, x_143, x_155, x_157, x_98); @@ -7010,6 +7019,8 @@ l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__29 = _init_l_Lean_E lean_mark_persistent(l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__29); l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__30 = _init_l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__30(); lean_mark_persistent(l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__30); +l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__31 = _init_l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__31(); +lean_mark_persistent(l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__31); l_Lean_Elab_Deriving_BEq_mkMutualBlock___closed__1 = _init_l_Lean_Elab_Deriving_BEq_mkMutualBlock___closed__1(); lean_mark_persistent(l_Lean_Elab_Deriving_BEq_mkMutualBlock___closed__1); l_Lean_Elab_Deriving_BEq_mkMutualBlock___closed__2 = _init_l_Lean_Elab_Deriving_BEq_mkMutualBlock___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/Deriving/DecEq.c b/stage0/stdlib/Lean/Elab/Deriving/DecEq.c index 7fac9322fb89..427ac32f4e9b 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/DecEq.c +++ b/stage0/stdlib/Lean/Elab/Deriving/DecEq.c @@ -419,6 +419,7 @@ 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_mkAuxFunction___closed__25; 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; @@ -4928,7 +4929,7 @@ static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__8() _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("def", 3); +x_1 = lean_mk_string_from_bytes("definition", 10); return x_1; } } @@ -4948,23 +4949,31 @@ static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__10( _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("declId", 6); +x_1 = lean_mk_string_from_bytes("def", 3); return x_1; } } static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__11() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("declId", 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkAuxFunction___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_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_mkAuxFunction___closed__1; -x_4 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__10; +x_4 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___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_Elab_Deriving_DecEq_mkAuxFunction___closed__12() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__13() { _start: { lean_object* x_1; @@ -4972,19 +4981,19 @@ x_1 = lean_mk_string_from_bytes("optDeclSig", 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__13() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkAuxFunction___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_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_mkAuxFunction___closed__1; -x_4 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__12; +x_4 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___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_Elab_Deriving_DecEq_mkAuxFunction___closed__14() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__15() { _start: { lean_object* x_1; @@ -4992,19 +5001,19 @@ x_1 = lean_mk_string_from_bytes("declValSimple", 13); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__15() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkAuxFunction___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_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_mkAuxFunction___closed__1; -x_4 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__14; +x_4 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___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_Elab_Deriving_DecEq_mkAuxFunction___closed__16() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__17() { _start: { lean_object* x_1; @@ -5012,7 +5021,7 @@ x_1 = lean_mk_string_from_bytes("Termination", 11); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__17() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__18() { _start: { lean_object* x_1; @@ -5020,19 +5029,19 @@ x_1 = lean_mk_string_from_bytes("suffix", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__18() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__19() { _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_mkAuxFunction___closed__16; -x_4 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__17; +x_3 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__17; +x_4 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__18; 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_mkAuxFunction___closed__19() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__20() { _start: { lean_object* x_1; lean_object* x_2; @@ -5041,7 +5050,7 @@ x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__20() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -5051,46 +5060,46 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__21() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkAuxFunction___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_Deriving_DecEq_mkAuxFunction___closed__20; +x_2 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__21; 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_mkAuxFunction___closed__22() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__23() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__20; +x_1 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___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_Elab_Deriving_DecEq_mkAuxFunction___closed__23() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__24() { _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_mkAuxFunction___closed__22; +x_2 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__23; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__24() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__21; -x_2 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__23; +x_1 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__22; +x_2 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___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); @@ -5147,10 +5156,10 @@ 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_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__20; +x_106 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__21; x_107 = l_Lean_addMacroScope(x_105, x_106, x_100); -x_108 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__19; -x_109 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__24; +x_108 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__20; +x_109 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__25; lean_inc(x_99); x_110 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_110, 0, x_99); @@ -5233,13 +5242,13 @@ x_34 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__5; lean_inc_n(x_28, 5); lean_inc(x_22); x_35 = l_Lean_Syntax_node6(x_22, x_34, x_28, x_28, x_33, x_28, x_28, x_28); -x_36 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__8; +x_36 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__10; lean_inc(x_22); x_37 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_37, 0, x_22); lean_ctor_set(x_37, 1, x_36); x_38 = lean_mk_syntax_ident(x_2); -x_39 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__11; +x_39 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__12; lean_inc(x_28); lean_inc(x_22); x_40 = l_Lean_Syntax_node2(x_22, x_39, x_38, x_28); @@ -5259,7 +5268,7 @@ lean_inc(x_22); x_46 = l_Lean_Syntax_node2(x_22, x_45, x_44, x_18); lean_inc(x_22); x_47 = l_Lean_Syntax_node1(x_22, x_26, x_46); -x_48 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__13; +x_48 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__14; lean_inc(x_22); x_49 = l_Lean_Syntax_node2(x_22, x_48, x_42, x_47); x_50 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__16; @@ -5267,11 +5276,11 @@ lean_inc(x_22); 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_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__18; +x_52 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__19; lean_inc_n(x_28, 2); lean_inc(x_22); x_53 = l_Lean_Syntax_node2(x_22, x_52, x_28, x_28); -x_54 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__15; +x_54 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__16; lean_inc(x_28); lean_inc(x_22); x_55 = l_Lean_Syntax_node4(x_22, x_54, x_51, x_15, x_53, x_28); @@ -5310,13 +5319,13 @@ x_69 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__5; lean_inc_n(x_63, 5); lean_inc(x_22); x_70 = l_Lean_Syntax_node6(x_22, x_69, x_63, x_63, x_68, x_63, x_63, x_63); -x_71 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__8; +x_71 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__10; lean_inc(x_22); x_72 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_72, 0, x_22); lean_ctor_set(x_72, 1, x_71); x_73 = lean_mk_syntax_ident(x_2); -x_74 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__11; +x_74 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__12; lean_inc(x_63); lean_inc(x_22); x_75 = l_Lean_Syntax_node2(x_22, x_74, x_73, x_63); @@ -5336,7 +5345,7 @@ lean_inc(x_22); x_81 = l_Lean_Syntax_node2(x_22, x_80, x_79, x_18); lean_inc(x_22); x_82 = l_Lean_Syntax_node1(x_22, x_61, x_81); -x_83 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__13; +x_83 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__14; lean_inc(x_22); x_84 = l_Lean_Syntax_node2(x_22, x_83, x_77, x_82); x_85 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__16; @@ -5344,11 +5353,11 @@ lean_inc(x_22); x_86 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_86, 0, x_22); lean_ctor_set(x_86, 1, x_85); -x_87 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__18; +x_87 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__19; lean_inc_n(x_63, 2); lean_inc(x_22); x_88 = l_Lean_Syntax_node2(x_22, x_87, x_63, x_63); -x_89 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__15; +x_89 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__16; lean_inc(x_63); lean_inc(x_22); x_90 = l_Lean_Syntax_node4(x_22, x_89, x_86, x_15, x_88, x_63); @@ -9241,11 +9250,11 @@ lean_inc(x_21); x_208 = l_Lean_Syntax_node4(x_21, x_200, x_65, x_30, x_67, x_207); lean_inc(x_21); x_209 = l_Lean_Syntax_node2(x_21, x_202, x_56, x_208); -x_210 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__18; +x_210 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__19; lean_inc_n(x_30, 2); lean_inc(x_21); x_211 = l_Lean_Syntax_node2(x_21, x_210, x_30, x_30); -x_212 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__15; +x_212 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__16; lean_inc(x_30); lean_inc(x_21); x_213 = l_Lean_Syntax_node4(x_21, x_212, x_54, x_209, x_211, x_30); @@ -10687,6 +10696,8 @@ l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__23 = _init_l_Lean_Elab_Derivi lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__23); l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__24 = _init_l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__24(); lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__24); +l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__25 = _init_l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__25(); +lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__25); l_Lean_Elab_Deriving_DecEq_mkAuxFunctions___closed__1 = _init_l_Lean_Elab_Deriving_DecEq_mkAuxFunctions___closed__1(); lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_mkAuxFunctions___closed__1); l_Lean_Elab_Deriving_DecEq_mkAuxFunctions___closed__2 = _init_l_Lean_Elab_Deriving_DecEq_mkAuxFunctions___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/Deriving/FromToJson.c b/stage0/stdlib/Lean/Elab/Deriving/FromToJson.c index e2f68387ef90..b33b12d6c655 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/FromToJson.c +++ b/stage0/stdlib/Lean/Elab/Deriving/FromToJson.c @@ -210,6 +210,7 @@ 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_mkToJsonInstance_mkAlts___spec__5___lambda__1___closed__8; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_quoteNameMk(lean_object*); +static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__38; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstance___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Command_liftTermElabM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstance_mkAlts___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*); @@ -4493,7 +4494,7 @@ static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lam _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("def", 3); +x_1 = lean_mk_string_from_bytes("definition", 10); return x_1; } } @@ -4513,23 +4514,31 @@ static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lam _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("declId", 6); +x_1 = lean_mk_string_from_bytes("def", 3); return x_1; } } static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__17() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("declId", 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___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_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstance_mkAlts___spec__3___closed__1; x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstance_mkAlts___spec__3___closed__2; x_3 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__7; -x_4 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__16; +x_4 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___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_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__18() { +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__19() { _start: { lean_object* x_1; @@ -4537,19 +4546,19 @@ x_1 = lean_mk_string_from_bytes("optDeclSig", 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__19() { +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__20() { _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_Lean_Elab_Deriving_FromToJson_mkToJsonInstance_mkAlts___spec__3___closed__1; x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstance_mkAlts___spec__3___closed__2; x_3 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__7; -x_4 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__18; +x_4 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___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_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__20() { +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__21() { _start: { lean_object* x_1; @@ -4557,19 +4566,19 @@ x_1 = lean_mk_string_from_bytes("typeSpec", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__21() { +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__22() { _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_Lean_Elab_Deriving_FromToJson_mkToJsonInstance_mkAlts___spec__3___closed__1; x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstance_mkAlts___spec__3___closed__2; x_3 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstance_mkAlts___spec__3___closed__3; -x_4 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__20; +x_4 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__21; 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_FromToJson_mkToJsonInstance___lambda__3___closed__22() { +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__23() { _start: { lean_object* x_1; @@ -4577,7 +4586,7 @@ x_1 = lean_mk_string_from_bytes(":", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__23() { +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__24() { _start: { lean_object* x_1; lean_object* x_2; @@ -4586,7 +4595,7 @@ x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__24() { +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -4596,7 +4605,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__25() { +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__26() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -4606,53 +4615,53 @@ x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__26() { +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__27() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__25; +x_2 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__26; 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_FromToJson_mkToJsonInstance___lambda__3___closed__27() { +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__28() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__25; +x_1 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__26; 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_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__28() { +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__29() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__27; +x_2 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__28; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__29() { +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__30() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__26; -x_2 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__28; +x_1 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__27; +x_2 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__29; 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_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__30() { +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__31() { _start: { lean_object* x_1; @@ -4660,19 +4669,19 @@ x_1 = lean_mk_string_from_bytes("declValSimple", 13); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__31() { +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___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_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstance_mkAlts___spec__3___closed__1; x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstance_mkAlts___spec__3___closed__2; x_3 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__7; -x_4 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__30; +x_4 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___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_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__32() { +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__33() { _start: { lean_object* x_1; @@ -4680,7 +4689,7 @@ x_1 = lean_mk_string_from_bytes(":=", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__33() { +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__34() { _start: { lean_object* x_1; @@ -4688,7 +4697,7 @@ x_1 = lean_mk_string_from_bytes("Termination", 11); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__34() { +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__35() { _start: { lean_object* x_1; @@ -4696,19 +4705,19 @@ x_1 = lean_mk_string_from_bytes("suffix", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__35() { +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___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_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstance_mkAlts___spec__3___closed__1; x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstance_mkAlts___spec__3___closed__2; -x_3 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__33; -x_4 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__34; +x_3 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__34; +x_4 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___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_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__36() { +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__37() { _start: { lean_object* x_1; @@ -4716,14 +4725,14 @@ x_1 = lean_mk_string_from_bytes("partial", 7); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__37() { +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___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_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstance_mkAlts___spec__3___closed__1; x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstance_mkAlts___spec__3___closed__2; x_3 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__7; -x_4 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__36; +x_4 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__37; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } @@ -4929,12 +4938,12 @@ x_74 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__11 lean_inc_n(x_45, 5); lean_inc(x_38); x_75 = l_Lean_Syntax_node6(x_38, x_74, x_45, x_45, x_73, x_45, x_45, x_45); -x_76 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__14; +x_76 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__16; lean_inc(x_38); x_77 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_77, 0, x_38); lean_ctor_set(x_77, 1, x_76); -x_78 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__17; +x_78 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__18; lean_inc(x_45); lean_inc(x_38); x_79 = l_Lean_Syntax_node2(x_38, x_78, x_22, x_45); @@ -4947,39 +4956,39 @@ x_82 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_82, 0, x_38); lean_ctor_set(x_82, 1, x_43); lean_ctor_set(x_82, 2, x_81); -x_83 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__22; +x_83 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__23; lean_inc(x_38); x_84 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_84, 0, x_38); lean_ctor_set(x_84, 1, x_83); -x_85 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__24; +x_85 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__25; x_86 = l_Lean_addMacroScope(x_68, x_85, x_63); -x_87 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__23; -x_88 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__29; +x_87 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__24; +x_88 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__30; lean_inc(x_38); x_89 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_89, 0, x_38); lean_ctor_set(x_89, 1, x_87); lean_ctor_set(x_89, 2, x_86); lean_ctor_set(x_89, 3, x_88); -x_90 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__21; +x_90 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__22; lean_inc(x_38); x_91 = l_Lean_Syntax_node2(x_38, x_90, x_84, x_89); lean_inc(x_38); x_92 = l_Lean_Syntax_node1(x_38, x_43, x_91); -x_93 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__19; +x_93 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__20; lean_inc(x_38); x_94 = l_Lean_Syntax_node2(x_38, x_93, x_82, x_92); -x_95 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__32; +x_95 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__33; lean_inc(x_38); x_96 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_96, 0, x_38); lean_ctor_set(x_96, 1, x_95); -x_97 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__35; +x_97 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__36; lean_inc_n(x_45, 2); lean_inc(x_38); x_98 = l_Lean_Syntax_node2(x_38, x_97, x_45, x_45); -x_99 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__31; +x_99 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__32; lean_inc(x_45); lean_inc(x_38); x_100 = l_Lean_Syntax_node4(x_38, x_99, x_96, x_61, x_98, x_45); @@ -5047,12 +5056,12 @@ lean_inc(x_38); x_122 = l_Lean_Syntax_node1(x_38, x_121, x_120); lean_inc(x_38); x_123 = l_Lean_Syntax_node1(x_38, x_43, x_122); -x_124 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__36; +x_124 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__37; lean_inc(x_38); x_125 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_125, 0, x_38); lean_ctor_set(x_125, 1, x_124); -x_126 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__37; +x_126 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__38; lean_inc(x_38); x_127 = l_Lean_Syntax_node1(x_38, x_126, x_125); lean_inc(x_38); @@ -5061,12 +5070,12 @@ x_129 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__1 lean_inc_n(x_45, 4); lean_inc(x_38); x_130 = l_Lean_Syntax_node6(x_38, x_129, x_45, x_45, x_123, x_45, x_45, x_128); -x_131 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__14; +x_131 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__16; 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_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__17; +x_133 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__18; lean_inc(x_45); lean_inc(x_38); x_134 = l_Lean_Syntax_node2(x_38, x_133, x_22, x_45); @@ -5079,39 +5088,39 @@ x_137 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_137, 0, x_38); lean_ctor_set(x_137, 1, x_43); lean_ctor_set(x_137, 2, x_136); -x_138 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__22; +x_138 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__23; lean_inc(x_38); x_139 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_139, 0, x_38); lean_ctor_set(x_139, 1, x_138); -x_140 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__24; +x_140 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__25; x_141 = l_Lean_addMacroScope(x_118, x_140, x_113); -x_142 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__23; -x_143 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__29; +x_142 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__24; +x_143 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__30; lean_inc(x_38); x_144 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_144, 0, x_38); lean_ctor_set(x_144, 1, x_142); lean_ctor_set(x_144, 2, x_141); lean_ctor_set(x_144, 3, x_143); -x_145 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__21; +x_145 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__22; lean_inc(x_38); x_146 = l_Lean_Syntax_node2(x_38, x_145, x_139, x_144); lean_inc(x_38); x_147 = l_Lean_Syntax_node1(x_38, x_43, x_146); -x_148 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__19; +x_148 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__20; lean_inc(x_38); x_149 = l_Lean_Syntax_node2(x_38, x_148, x_137, x_147); -x_150 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__32; +x_150 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__33; lean_inc(x_38); x_151 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_151, 0, x_38); lean_ctor_set(x_151, 1, x_150); -x_152 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__35; +x_152 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__36; lean_inc_n(x_45, 2); lean_inc(x_38); x_153 = l_Lean_Syntax_node2(x_38, x_152, x_45, x_45); -x_154 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__31; +x_154 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__32; lean_inc(x_45); lean_inc(x_38); x_155 = l_Lean_Syntax_node4(x_38, x_154, x_151, x_111, x_153, x_45); @@ -5527,7 +5536,7 @@ x_47 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__11 lean_inc_n(x_41, 5); lean_inc(x_32); x_48 = l_Lean_Syntax_node6(x_32, x_47, x_41, x_41, x_46, x_41, x_41, x_41); -x_49 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__14; +x_49 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__16; lean_inc(x_32); x_50 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_50, 0, x_32); @@ -5547,32 +5556,32 @@ x_57 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_57, 0, x_32); lean_ctor_set(x_57, 1, x_39); lean_ctor_set(x_57, 2, x_56); -x_58 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__22; +x_58 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__23; lean_inc(x_32); x_59 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_59, 0, x_32); lean_ctor_set(x_59, 1, x_58); -x_60 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__24; +x_60 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__25; lean_inc(x_33); lean_inc(x_38); x_61 = l_Lean_addMacroScope(x_38, x_60, x_33); -x_62 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__23; -x_63 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__29; +x_62 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__24; +x_63 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__30; lean_inc(x_32); x_64 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_64, 0, x_32); lean_ctor_set(x_64, 1, x_62); lean_ctor_set(x_64, 2, x_61); lean_ctor_set(x_64, 3, x_63); -x_65 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__21; +x_65 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__22; lean_inc(x_32); x_66 = l_Lean_Syntax_node2(x_32, x_65, x_59, x_64); lean_inc(x_32); x_67 = l_Lean_Syntax_node1(x_32, x_39, x_66); -x_68 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__19; +x_68 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__20; lean_inc(x_32); x_69 = l_Lean_Syntax_node2(x_32, x_68, x_57, x_67); -x_70 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__32; +x_70 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__33; lean_inc(x_32); x_71 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_71, 0, x_32); @@ -5634,11 +5643,11 @@ x_96 = l_Lean_Syntax_node2(x_32, x_95, x_83, x_94); x_97 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__4___closed__2; lean_inc(x_32); x_98 = l_Lean_Syntax_node3(x_32, x_97, x_76, x_78, x_96); -x_99 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__35; +x_99 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__36; lean_inc_n(x_41, 2); lean_inc(x_32); x_100 = l_Lean_Syntax_node2(x_32, x_99, x_41, x_41); -x_101 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__31; +x_101 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__32; lean_inc(x_41); lean_inc(x_32); x_102 = l_Lean_Syntax_node4(x_32, x_101, x_71, x_98, x_100, x_41); @@ -5649,7 +5658,7 @@ lean_dec(x_51); x_103 = l_Lean_instInhabitedName; x_104 = l___private_Init_Util_0__outOfBounds___rarg(x_103); x_105 = lean_mk_syntax_ident(x_104); -x_106 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__17; +x_106 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__18; lean_inc(x_41); lean_inc(x_32); x_107 = l_Lean_Syntax_node2(x_32, x_106, x_105, x_41); @@ -5723,7 +5732,7 @@ lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; x_129 = lean_array_fget(x_51, x_53); lean_dec(x_51); x_130 = lean_mk_syntax_ident(x_129); -x_131 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__17; +x_131 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__18; lean_inc(x_41); lean_inc(x_32); x_132 = l_Lean_Syntax_node2(x_32, x_131, x_130, x_41); @@ -10228,12 +10237,12 @@ x_138 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__1 lean_inc_n(x_132, 5); lean_inc(x_124); x_139 = l_Lean_Syntax_node6(x_124, x_138, x_132, x_132, x_137, x_132, x_132, x_132); -x_140 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__14; +x_140 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__16; lean_inc(x_124); x_141 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_141, 0, x_124); lean_ctor_set(x_141, 1, x_140); -x_142 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__17; +x_142 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__18; lean_inc(x_132); lean_inc(x_124); x_143 = l_Lean_Syntax_node2(x_124, x_142, x_7, x_132); @@ -10265,21 +10274,21 @@ lean_ctor_set(x_153, 3, x_151); lean_inc(x_6); lean_inc(x_124); x_154 = l_Lean_Syntax_node1(x_124, x_6, x_153); -x_155 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__22; +x_155 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__23; lean_inc(x_124); x_156 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_156, 0, x_124); lean_ctor_set(x_156, 1, x_155); -x_157 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__24; +x_157 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__25; lean_inc(x_125); lean_inc(x_130); x_158 = l_Lean_addMacroScope(x_130, x_157, x_125); -x_159 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__25; +x_159 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__26; lean_inc(x_8); x_160 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_160, 0, x_159); lean_ctor_set(x_160, 1, x_8); -x_161 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__27; +x_161 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__28; lean_inc(x_8); x_162 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_162, 0, x_161); @@ -10287,7 +10296,7 @@ lean_ctor_set(x_162, 1, x_8); x_163 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_163, 0, x_160); lean_ctor_set(x_163, 1, x_162); -x_164 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__23; +x_164 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__24; lean_inc(x_124); x_165 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_165, 0, x_124); @@ -10362,24 +10371,24 @@ lean_inc(x_124); x_189 = l_Lean_Syntax_node2(x_124, x_6, x_188, x_120); lean_inc(x_124); x_190 = l_Lean_Syntax_node2(x_124, x_9, x_180, x_189); -x_191 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__21; +x_191 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__22; lean_inc(x_124); x_192 = l_Lean_Syntax_node2(x_124, x_191, x_156, x_190); lean_inc(x_124); x_193 = l_Lean_Syntax_node1(x_124, x_6, x_192); -x_194 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__19; +x_194 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__20; lean_inc(x_124); x_195 = l_Lean_Syntax_node2(x_124, x_194, x_172, x_193); -x_196 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__32; +x_196 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__33; lean_inc(x_124); x_197 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_197, 0, x_124); lean_ctor_set(x_197, 1, x_196); -x_198 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__35; +x_198 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__36; lean_inc_n(x_132, 2); lean_inc(x_124); x_199 = l_Lean_Syntax_node2(x_124, x_198, x_132, x_132); -x_200 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__31; +x_200 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__32; lean_inc(x_132); lean_inc(x_124); x_201 = l_Lean_Syntax_node4(x_124, x_200, x_197, x_11, x_199, x_132); @@ -10453,12 +10462,12 @@ x_39 = l_Lean_Syntax_node1(x_27, x_38, x_37); lean_inc(x_6); lean_inc(x_27); x_40 = l_Lean_Syntax_node1(x_27, x_6, x_39); -x_41 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__36; +x_41 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__37; lean_inc(x_27); x_42 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_42, 0, x_27); lean_ctor_set(x_42, 1, x_41); -x_43 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__37; +x_43 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__38; lean_inc(x_27); x_44 = l_Lean_Syntax_node1(x_27, x_43, x_42); lean_inc(x_6); @@ -10468,12 +10477,12 @@ x_46 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__11 lean_inc_n(x_35, 4); lean_inc(x_27); x_47 = l_Lean_Syntax_node6(x_27, x_46, x_35, x_35, x_40, x_35, x_35, x_45); -x_48 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__14; +x_48 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__16; lean_inc(x_27); x_49 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_49, 0, x_27); lean_ctor_set(x_49, 1, x_48); -x_50 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__17; +x_50 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__18; lean_inc(x_35); lean_inc(x_27); x_51 = l_Lean_Syntax_node2(x_27, x_50, x_7, x_35); @@ -10505,21 +10514,21 @@ lean_ctor_set(x_61, 3, x_59); lean_inc(x_6); lean_inc(x_27); x_62 = l_Lean_Syntax_node1(x_27, x_6, x_61); -x_63 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__22; +x_63 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__23; lean_inc(x_27); x_64 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_64, 0, x_27); lean_ctor_set(x_64, 1, x_63); -x_65 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__24; +x_65 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__25; lean_inc(x_28); lean_inc(x_33); x_66 = l_Lean_addMacroScope(x_33, x_65, x_28); -x_67 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__25; +x_67 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__26; lean_inc(x_8); x_68 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_68, 0, x_67); lean_ctor_set(x_68, 1, x_8); -x_69 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__27; +x_69 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__28; lean_inc(x_8); x_70 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_70, 0, x_69); @@ -10527,7 +10536,7 @@ lean_ctor_set(x_70, 1, x_8); x_71 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_71, 0, x_68); lean_ctor_set(x_71, 1, x_70); -x_72 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__23; +x_72 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__24; lean_inc(x_27); x_73 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_73, 0, x_27); @@ -10602,24 +10611,24 @@ lean_inc(x_27); x_97 = l_Lean_Syntax_node2(x_27, x_6, x_96, x_23); lean_inc(x_27); x_98 = l_Lean_Syntax_node2(x_27, x_9, x_88, x_97); -x_99 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__21; +x_99 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__22; lean_inc(x_27); x_100 = l_Lean_Syntax_node2(x_27, x_99, x_64, x_98); lean_inc(x_27); x_101 = l_Lean_Syntax_node1(x_27, x_6, x_100); -x_102 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__19; +x_102 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__20; lean_inc(x_27); x_103 = l_Lean_Syntax_node2(x_27, x_102, x_80, x_101); -x_104 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__32; +x_104 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__33; lean_inc(x_27); x_105 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_105, 0, x_27); lean_ctor_set(x_105, 1, x_104); -x_106 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__35; +x_106 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__36; lean_inc_n(x_35, 2); lean_inc(x_27); x_107 = l_Lean_Syntax_node2(x_27, x_106, x_35, x_35); -x_108 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__31; +x_108 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__32; lean_inc(x_35); lean_inc(x_27); x_109 = l_Lean_Syntax_node4(x_27, x_108, x_105, x_11, x_107, x_35); @@ -11151,7 +11160,7 @@ static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstance___l { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__25; +x_2 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__26; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); @@ -11163,7 +11172,7 @@ static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstance___l { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstance___lambda__3___closed__1; -x_2 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__28; +x_2 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__29; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -11437,7 +11446,7 @@ x_52 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__11 lean_inc_n(x_46, 5); lean_inc(x_37); x_53 = l_Lean_Syntax_node6(x_37, x_52, x_46, x_46, x_51, x_46, x_46, x_46); -x_54 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__14; +x_54 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__16; lean_inc(x_37); x_55 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_55, 0, x_37); @@ -11470,16 +11479,16 @@ lean_ctor_set(x_67, 2, x_64); lean_ctor_set(x_67, 3, x_65); lean_inc(x_37); x_68 = l_Lean_Syntax_node1(x_37, x_44, x_67); -x_69 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__22; +x_69 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__23; lean_inc(x_37); x_70 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_70, 0, x_37); lean_ctor_set(x_70, 1, x_69); -x_71 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__24; +x_71 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__25; lean_inc(x_38); lean_inc(x_43); x_72 = l_Lean_addMacroScope(x_43, x_71, x_38); -x_73 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__23; +x_73 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__24; x_74 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstance___lambda__3___closed__2; lean_inc(x_37); x_75 = lean_alloc_ctor(3, 4, 0); @@ -11532,15 +11541,15 @@ x_93 = l_Lean_Syntax_node2(x_37, x_44, x_92, x_34); x_94 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstance_mkAlts___spec__5___lambda__1___closed__3; lean_inc(x_37); x_95 = l_Lean_Syntax_node2(x_37, x_94, x_87, x_93); -x_96 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__21; +x_96 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__22; lean_inc(x_37); x_97 = l_Lean_Syntax_node2(x_37, x_96, x_70, x_95); lean_inc(x_37); x_98 = l_Lean_Syntax_node1(x_37, x_44, x_97); -x_99 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__19; +x_99 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__20; lean_inc(x_37); x_100 = l_Lean_Syntax_node2(x_37, x_99, x_82, x_98); -x_101 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__32; +x_101 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__33; lean_inc(x_37); x_102 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_102, 0, x_37); @@ -11621,11 +11630,11 @@ x_136 = l_Lean_Syntax_node1(x_37, x_135, x_134); x_137 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstance_mkAlts___spec__7___lambda__1___closed__31; lean_inc(x_37); x_138 = l_Lean_Syntax_node2(x_37, x_137, x_104, x_136); -x_139 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__35; +x_139 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__36; lean_inc_n(x_46, 2); lean_inc(x_37); x_140 = l_Lean_Syntax_node2(x_37, x_139, x_46, x_46); -x_141 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__31; +x_141 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__32; lean_inc(x_46); lean_inc(x_37); x_142 = l_Lean_Syntax_node4(x_37, x_141, x_102, x_138, x_140, x_46); @@ -11636,7 +11645,7 @@ lean_dec(x_56); x_143 = l_Lean_instInhabitedName; x_144 = l___private_Init_Util_0__outOfBounds___rarg(x_143); x_145 = lean_mk_syntax_ident(x_144); -x_146 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__17; +x_146 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__18; lean_inc(x_46); lean_inc(x_37); x_147 = l_Lean_Syntax_node2(x_37, x_146, x_145, x_46); @@ -11710,7 +11719,7 @@ lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; x_169 = lean_array_fget(x_56, x_15); lean_dec(x_56); x_170 = lean_mk_syntax_ident(x_169); -x_171 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__17; +x_171 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__18; lean_inc(x_46); lean_inc(x_37); x_172 = l_Lean_Syntax_node2(x_37, x_171, x_170, x_46); @@ -13140,6 +13149,8 @@ l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__36 = _ini lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__36); l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__37 = _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__37(); lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__37); +l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__38 = _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__38(); +lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__38); l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__4___closed__1 = _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__4___closed__1(); lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__4___closed__1); l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__4___closed__2 = _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__4___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/Deriving/Hashable.c b/stage0/stdlib/Lean/Elab/Deriving/Hashable.c index f075ac16b8cb..1ca072118d1d 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/Hashable.c +++ b/stage0/stdlib/Lean/Elab/Deriving/Hashable.c @@ -25,6 +25,7 @@ static lean_object* l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___cl static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__7; static lean_object* l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__30; lean_object* lean_whnf(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__32; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__19; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__18; static lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__1___closed__1; @@ -2553,7 +2554,7 @@ static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__ _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("def", 3); +x_1 = lean_mk_string_from_bytes("definition", 10); return x_1; } } @@ -2573,23 +2574,31 @@ static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__ _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("declId", 6); +x_1 = lean_mk_string_from_bytes("def", 3); return x_1; } } static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__11() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("declId", 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___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_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__3___closed__1; x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__3___closed__2; x_3 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__1; -x_4 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__10; +x_4 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___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_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__12() { +static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__13() { _start: { lean_object* x_1; @@ -2597,19 +2606,19 @@ x_1 = lean_mk_string_from_bytes("optDeclSig", 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__13() { +static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___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_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__3___closed__1; x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__3___closed__2; x_3 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__1; -x_4 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__12; +x_4 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___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_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__14() { +static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__15() { _start: { lean_object* x_1; @@ -2617,19 +2626,19 @@ x_1 = lean_mk_string_from_bytes("typeSpec", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__15() { +static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___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_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__3___closed__1; x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__3___closed__2; x_3 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__3___closed__3; -x_4 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__14; +x_4 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___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_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__16() { +static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__17() { _start: { lean_object* x_1; @@ -2637,7 +2646,7 @@ x_1 = lean_mk_string_from_bytes(":", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__17() { +static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__18() { _start: { lean_object* x_1; @@ -2645,72 +2654,72 @@ x_1 = lean_mk_string_from_bytes("UInt64", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__18() { +static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__19() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__17; +x_1 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__18; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__19() { +static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___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_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__17; +x_2 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__18; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__20() { +static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__19; +x_2 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__20; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__21() { +static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__22() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__19; +x_1 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___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_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__22() { +static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__21; +x_2 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__22; 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_Hashable_mkAuxFunction___lambda__1___closed__23() { +static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__20; -x_2 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__22; +x_1 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__21; +x_2 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___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_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__24() { +static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__25() { _start: { lean_object* x_1; @@ -2718,19 +2727,19 @@ x_1 = lean_mk_string_from_bytes("declValSimple", 13); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__25() { +static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___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_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__3___closed__1; x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__3___closed__2; x_3 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__1; -x_4 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__24; +x_4 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___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_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__26() { +static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__27() { _start: { lean_object* x_1; @@ -2738,7 +2747,7 @@ x_1 = lean_mk_string_from_bytes(":=", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__27() { +static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__28() { _start: { lean_object* x_1; @@ -2746,7 +2755,7 @@ x_1 = lean_mk_string_from_bytes("Termination", 11); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__28() { +static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__29() { _start: { lean_object* x_1; @@ -2754,19 +2763,19 @@ x_1 = lean_mk_string_from_bytes("suffix", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__29() { +static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__30() { _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_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__3___closed__1; x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__3___closed__2; -x_3 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__27; -x_4 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__28; +x_3 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__28; +x_4 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__29; 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_Hashable_mkAuxFunction___lambda__1___closed__30() { +static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__31() { _start: { lean_object* x_1; @@ -2774,14 +2783,14 @@ x_1 = lean_mk_string_from_bytes("partial", 7); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__31() { +static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___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_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__3___closed__1; x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__3___closed__2; x_3 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__1; -x_4 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__30; +x_4 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__31; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } @@ -2835,13 +2844,13 @@ x_32 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__5; lean_inc_n(x_26, 5); lean_inc(x_17); x_33 = l_Lean_Syntax_node6(x_17, x_32, x_26, x_26, x_31, x_26, x_26, x_26); -x_34 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__8; +x_34 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__10; lean_inc(x_17); x_35 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_35, 0, x_17); lean_ctor_set(x_35, 1, x_34); x_36 = lean_mk_syntax_ident(x_3); -x_37 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__11; +x_37 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__12; lean_inc(x_26); lean_inc(x_17); x_38 = l_Lean_Syntax_node2(x_17, x_37, x_36, x_26); @@ -2851,39 +2860,39 @@ x_40 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_40, 0, x_17); lean_ctor_set(x_40, 1, x_24); lean_ctor_set(x_40, 2, x_39); -x_41 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__16; +x_41 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__17; lean_inc(x_17); x_42 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_42, 0, x_17); lean_ctor_set(x_42, 1, x_41); -x_43 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__19; +x_43 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__20; x_44 = l_Lean_addMacroScope(x_23, x_43, x_18); -x_45 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__18; -x_46 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__23; +x_45 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__19; +x_46 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__24; lean_inc(x_17); x_47 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_47, 0, x_17); lean_ctor_set(x_47, 1, x_45); lean_ctor_set(x_47, 2, x_44); lean_ctor_set(x_47, 3, x_46); -x_48 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__15; +x_48 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__16; lean_inc(x_17); x_49 = l_Lean_Syntax_node2(x_17, x_48, x_42, x_47); lean_inc(x_17); x_50 = l_Lean_Syntax_node1(x_17, x_24, x_49); -x_51 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__13; +x_51 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__14; lean_inc(x_17); x_52 = l_Lean_Syntax_node2(x_17, x_51, x_40, x_50); -x_53 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__26; +x_53 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__27; lean_inc(x_17); x_54 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_54, 0, x_17); lean_ctor_set(x_54, 1, x_53); -x_55 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__29; +x_55 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__30; lean_inc_n(x_26, 2); lean_inc(x_17); x_56 = l_Lean_Syntax_node2(x_17, x_55, x_26, x_26); -x_57 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__25; +x_57 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__26; lean_inc(x_26); lean_inc(x_17); x_58 = l_Lean_Syntax_node4(x_17, x_57, x_54, x_4, x_56, x_26); @@ -2928,13 +2937,13 @@ x_75 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__5; lean_inc_n(x_69, 5); lean_inc(x_17); x_76 = l_Lean_Syntax_node6(x_17, x_75, x_69, x_69, x_74, x_69, x_69, x_69); -x_77 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__8; +x_77 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__10; lean_inc(x_17); x_78 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_78, 0, x_17); lean_ctor_set(x_78, 1, x_77); x_79 = lean_mk_syntax_ident(x_3); -x_80 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__11; +x_80 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__12; lean_inc(x_69); lean_inc(x_17); x_81 = l_Lean_Syntax_node2(x_17, x_80, x_79, x_69); @@ -2944,39 +2953,39 @@ x_83 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_83, 0, x_17); lean_ctor_set(x_83, 1, x_67); lean_ctor_set(x_83, 2, x_82); -x_84 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__16; +x_84 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__17; lean_inc(x_17); x_85 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_85, 0, x_17); lean_ctor_set(x_85, 1, x_84); -x_86 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__19; +x_86 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__20; x_87 = l_Lean_addMacroScope(x_66, x_86, x_18); -x_88 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__18; -x_89 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__23; +x_88 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__19; +x_89 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__24; lean_inc(x_17); x_90 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_90, 0, x_17); lean_ctor_set(x_90, 1, x_88); lean_ctor_set(x_90, 2, x_87); lean_ctor_set(x_90, 3, x_89); -x_91 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__15; +x_91 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__16; lean_inc(x_17); x_92 = l_Lean_Syntax_node2(x_17, x_91, x_85, x_90); lean_inc(x_17); x_93 = l_Lean_Syntax_node1(x_17, x_67, x_92); -x_94 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__13; +x_94 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__14; lean_inc(x_17); x_95 = l_Lean_Syntax_node2(x_17, x_94, x_83, x_93); -x_96 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__26; +x_96 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__27; lean_inc(x_17); x_97 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_97, 0, x_17); lean_ctor_set(x_97, 1, x_96); -x_98 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__29; +x_98 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__30; lean_inc_n(x_69, 2); lean_inc(x_17); x_99 = l_Lean_Syntax_node2(x_17, x_98, x_69, x_69); -x_100 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__25; +x_100 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__26; lean_inc(x_69); lean_inc(x_17); x_101 = l_Lean_Syntax_node4(x_17, x_100, x_97, x_4, x_99, x_69); @@ -3031,12 +3040,12 @@ lean_inc(x_110); x_123 = l_Lean_Syntax_node1(x_110, x_122, x_121); lean_inc(x_110); x_124 = l_Lean_Syntax_node1(x_110, x_117, x_123); -x_125 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__30; +x_125 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__31; lean_inc(x_110); x_126 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_126, 0, x_110); lean_ctor_set(x_126, 1, x_125); -x_127 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__31; +x_127 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__32; lean_inc(x_110); x_128 = l_Lean_Syntax_node1(x_110, x_127, x_126); lean_inc(x_110); @@ -3045,13 +3054,13 @@ x_130 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__5; lean_inc_n(x_119, 4); lean_inc(x_110); x_131 = l_Lean_Syntax_node6(x_110, x_130, x_119, x_119, x_124, x_119, x_119, x_129); -x_132 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__8; +x_132 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__10; lean_inc(x_110); x_133 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_133, 0, x_110); lean_ctor_set(x_133, 1, x_132); x_134 = lean_mk_syntax_ident(x_3); -x_135 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__11; +x_135 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__12; lean_inc(x_119); lean_inc(x_110); x_136 = l_Lean_Syntax_node2(x_110, x_135, x_134, x_119); @@ -3061,39 +3070,39 @@ x_138 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_138, 0, x_110); lean_ctor_set(x_138, 1, x_117); lean_ctor_set(x_138, 2, x_137); -x_139 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__16; +x_139 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__17; lean_inc(x_110); x_140 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_140, 0, x_110); lean_ctor_set(x_140, 1, x_139); -x_141 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__19; +x_141 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__20; x_142 = l_Lean_addMacroScope(x_116, x_141, x_111); -x_143 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__18; -x_144 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__23; +x_143 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__19; +x_144 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__24; lean_inc(x_110); x_145 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_145, 0, x_110); lean_ctor_set(x_145, 1, x_143); lean_ctor_set(x_145, 2, x_142); lean_ctor_set(x_145, 3, x_144); -x_146 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__15; +x_146 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__16; lean_inc(x_110); x_147 = l_Lean_Syntax_node2(x_110, x_146, x_140, x_145); lean_inc(x_110); x_148 = l_Lean_Syntax_node1(x_110, x_117, x_147); -x_149 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__13; +x_149 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__14; lean_inc(x_110); x_150 = l_Lean_Syntax_node2(x_110, x_149, x_138, x_148); -x_151 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__26; +x_151 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__27; lean_inc(x_110); x_152 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_152, 0, x_110); lean_ctor_set(x_152, 1, x_151); -x_153 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__29; +x_153 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__30; lean_inc_n(x_119, 2); lean_inc(x_110); x_154 = l_Lean_Syntax_node2(x_110, x_153, x_119, x_119); -x_155 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__25; +x_155 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__26; lean_inc(x_119); lean_inc(x_110); x_156 = l_Lean_Syntax_node4(x_110, x_155, x_152, x_4, x_154, x_119); @@ -3134,12 +3143,12 @@ lean_inc(x_110); x_171 = l_Lean_Syntax_node1(x_110, x_170, x_169); lean_inc(x_110); x_172 = l_Lean_Syntax_node1(x_110, x_165, x_171); -x_173 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__30; +x_173 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__31; lean_inc(x_110); x_174 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_174, 0, x_110); lean_ctor_set(x_174, 1, x_173); -x_175 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__31; +x_175 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__32; lean_inc(x_110); x_176 = l_Lean_Syntax_node1(x_110, x_175, x_174); lean_inc(x_110); @@ -3148,13 +3157,13 @@ x_178 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__5; lean_inc_n(x_167, 4); lean_inc(x_110); x_179 = l_Lean_Syntax_node6(x_110, x_178, x_167, x_167, x_172, x_167, x_167, x_177); -x_180 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__8; +x_180 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__10; lean_inc(x_110); x_181 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_181, 0, x_110); lean_ctor_set(x_181, 1, x_180); x_182 = lean_mk_syntax_ident(x_3); -x_183 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__11; +x_183 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__12; lean_inc(x_167); lean_inc(x_110); x_184 = l_Lean_Syntax_node2(x_110, x_183, x_182, x_167); @@ -3164,39 +3173,39 @@ x_186 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_186, 0, x_110); lean_ctor_set(x_186, 1, x_165); lean_ctor_set(x_186, 2, x_185); -x_187 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__16; +x_187 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__17; lean_inc(x_110); x_188 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_188, 0, x_110); lean_ctor_set(x_188, 1, x_187); -x_189 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__19; +x_189 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__20; x_190 = l_Lean_addMacroScope(x_164, x_189, x_111); -x_191 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__18; -x_192 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__23; +x_191 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__19; +x_192 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__24; lean_inc(x_110); x_193 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_193, 0, x_110); lean_ctor_set(x_193, 1, x_191); lean_ctor_set(x_193, 2, x_190); lean_ctor_set(x_193, 3, x_192); -x_194 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__15; +x_194 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__16; lean_inc(x_110); x_195 = l_Lean_Syntax_node2(x_110, x_194, x_188, x_193); lean_inc(x_110); x_196 = l_Lean_Syntax_node1(x_110, x_165, x_195); -x_197 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__13; +x_197 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__14; lean_inc(x_110); x_198 = l_Lean_Syntax_node2(x_110, x_197, x_186, x_196); -x_199 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__26; +x_199 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__27; lean_inc(x_110); x_200 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_200, 0, x_110); lean_ctor_set(x_200, 1, x_199); -x_201 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__29; +x_201 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__30; lean_inc_n(x_167, 2); lean_inc(x_110); x_202 = l_Lean_Syntax_node2(x_110, x_201, x_167, x_167); -x_203 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__25; +x_203 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__26; lean_inc(x_167); lean_inc(x_110); x_204 = l_Lean_Syntax_node4(x_110, x_203, x_200, x_4, x_202, x_167); @@ -5076,6 +5085,8 @@ l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__30 = _init_l_L lean_mark_persistent(l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__30); l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__31 = _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__31(); lean_mark_persistent(l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__31); +l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__32 = _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__32(); +lean_mark_persistent(l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__32); l_Lean_Elab_Deriving_Hashable_mkHashFuncs___closed__1 = _init_l_Lean_Elab_Deriving_Hashable_mkHashFuncs___closed__1(); lean_mark_persistent(l_Lean_Elab_Deriving_Hashable_mkHashFuncs___closed__1); l_Lean_Elab_Deriving_Hashable_mkHashFuncs___closed__2 = _init_l_Lean_Elab_Deriving_Hashable_mkHashFuncs___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/Deriving/Ord.c b/stage0/stdlib/Lean/Elab/Deriving/Ord.c index 22ffb29f1231..ac225b31dd76 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/Ord.c +++ b/stage0/stdlib/Lean/Elab/Deriving/Ord.c @@ -171,6 +171,7 @@ 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; +static lean_object* l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__31; 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; @@ -3152,7 +3153,7 @@ static lean_object* _init_l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___c _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("def", 3); +x_1 = lean_mk_string_from_bytes("definition", 10); return x_1; } } @@ -3172,23 +3173,31 @@ static lean_object* _init_l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___c _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("declId", 6); +x_1 = lean_mk_string_from_bytes("def", 3); return x_1; } } static lean_object* _init_l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__13() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("declId", 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___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_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__3___closed__1; x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__3___closed__2; x_3 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__1; -x_4 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__12; +x_4 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___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_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__14() { +static lean_object* _init_l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__15() { _start: { lean_object* x_1; @@ -3196,19 +3205,19 @@ x_1 = lean_mk_string_from_bytes("optDeclSig", 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__15() { +static lean_object* _init_l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___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_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__3___closed__1; x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__3___closed__2; x_3 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__1; -x_4 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__14; +x_4 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___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_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__16() { +static lean_object* _init_l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__17() { _start: { lean_object* x_1; @@ -3216,19 +3225,19 @@ x_1 = lean_mk_string_from_bytes("typeSpec", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__17() { +static lean_object* _init_l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___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_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__3___closed__1; x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__3___closed__2; x_3 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__3___closed__3; -x_4 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__16; +x_4 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___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_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__18() { +static lean_object* _init_l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__19() { _start: { lean_object* x_1; @@ -3236,7 +3245,7 @@ x_1 = lean_mk_string_from_bytes(":", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__19() { +static lean_object* _init_l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__20() { _start: { lean_object* x_1; lean_object* x_2; @@ -3245,7 +3254,7 @@ x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__20() { +static lean_object* _init_l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -3255,53 +3264,53 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__21() { +static lean_object* _init_l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__20; +x_2 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__21; 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_Ord_mkAuxFunction___lambda__1___closed__22() { +static lean_object* _init_l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__23() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__20; +x_1 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___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_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__23() { +static lean_object* _init_l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__22; +x_2 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__23; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__24() { +static lean_object* _init_l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__21; -x_2 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__23; +x_1 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__22; +x_2 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___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_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__25() { +static lean_object* _init_l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__26() { _start: { lean_object* x_1; @@ -3309,19 +3318,19 @@ x_1 = lean_mk_string_from_bytes("declValSimple", 13); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__26() { +static lean_object* _init_l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__27() { _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_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__3___closed__1; x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__3___closed__2; x_3 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__1; -x_4 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__25; +x_4 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__26; 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_Ord_mkAuxFunction___lambda__1___closed__27() { +static lean_object* _init_l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__28() { _start: { lean_object* x_1; @@ -3329,7 +3338,7 @@ x_1 = lean_mk_string_from_bytes(":=", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__28() { +static lean_object* _init_l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__29() { _start: { lean_object* x_1; @@ -3337,7 +3346,7 @@ x_1 = lean_mk_string_from_bytes("Termination", 11); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__29() { +static lean_object* _init_l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__30() { _start: { lean_object* x_1; @@ -3345,14 +3354,14 @@ x_1 = lean_mk_string_from_bytes("suffix", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__30() { +static lean_object* _init_l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__31() { _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_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__3___closed__1; x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__3___closed__2; -x_3 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__28; -x_4 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__29; +x_3 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__29; +x_4 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__30; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } @@ -3410,13 +3419,13 @@ x_138 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__5; lean_inc_n(x_132, 5); lean_inc(x_123); x_139 = l_Lean_Syntax_node6(x_123, x_138, x_132, x_132, x_137, x_132, x_132, x_132); -x_140 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__10; +x_140 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__12; lean_inc(x_123); x_141 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_141, 0, x_123); lean_ctor_set(x_141, 1, x_140); x_142 = lean_mk_syntax_ident(x_2); -x_143 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__13; +x_143 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__14; lean_inc(x_132); lean_inc(x_123); x_144 = l_Lean_Syntax_node2(x_123, x_143, x_142, x_132); @@ -3426,39 +3435,39 @@ x_146 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_146, 0, x_123); lean_ctor_set(x_146, 1, x_130); lean_ctor_set(x_146, 2, x_145); -x_147 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__18; +x_147 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__19; lean_inc(x_123); x_148 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_148, 0, x_123); lean_ctor_set(x_148, 1, x_147); -x_149 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__20; +x_149 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__21; x_150 = l_Lean_addMacroScope(x_129, x_149, x_124); -x_151 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__19; -x_152 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__24; +x_151 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__20; +x_152 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__25; lean_inc(x_123); x_153 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_153, 0, x_123); lean_ctor_set(x_153, 1, x_151); lean_ctor_set(x_153, 2, x_150); lean_ctor_set(x_153, 3, x_152); -x_154 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__17; +x_154 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__18; lean_inc(x_123); x_155 = l_Lean_Syntax_node2(x_123, x_154, x_148, x_153); lean_inc(x_123); x_156 = l_Lean_Syntax_node1(x_123, x_130, x_155); -x_157 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__15; +x_157 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__16; lean_inc(x_123); x_158 = l_Lean_Syntax_node2(x_123, x_157, x_146, x_156); -x_159 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__27; +x_159 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__28; lean_inc(x_123); x_160 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_160, 0, x_123); lean_ctor_set(x_160, 1, x_159); -x_161 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__30; +x_161 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__31; lean_inc_n(x_132, 2); lean_inc(x_123); x_162 = l_Lean_Syntax_node2(x_123, x_161, x_132, x_132); -x_163 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__26; +x_163 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__27; lean_inc(x_132); lean_inc(x_123); x_164 = l_Lean_Syntax_node4(x_123, x_163, x_160, x_5, x_162, x_132); @@ -3503,13 +3512,13 @@ x_181 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__5; lean_inc_n(x_175, 5); lean_inc(x_123); x_182 = l_Lean_Syntax_node6(x_123, x_181, x_175, x_175, x_180, x_175, x_175, x_175); -x_183 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__10; +x_183 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__12; lean_inc(x_123); x_184 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_184, 0, x_123); lean_ctor_set(x_184, 1, x_183); x_185 = lean_mk_syntax_ident(x_2); -x_186 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__13; +x_186 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__14; lean_inc(x_175); lean_inc(x_123); x_187 = l_Lean_Syntax_node2(x_123, x_186, x_185, x_175); @@ -3519,39 +3528,39 @@ x_189 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_189, 0, x_123); lean_ctor_set(x_189, 1, x_173); lean_ctor_set(x_189, 2, x_188); -x_190 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__18; +x_190 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__19; lean_inc(x_123); x_191 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_191, 0, x_123); lean_ctor_set(x_191, 1, x_190); -x_192 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__20; +x_192 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__21; x_193 = l_Lean_addMacroScope(x_172, x_192, x_124); -x_194 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__19; -x_195 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__24; +x_194 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__20; +x_195 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__25; lean_inc(x_123); x_196 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_196, 0, x_123); lean_ctor_set(x_196, 1, x_194); lean_ctor_set(x_196, 2, x_193); lean_ctor_set(x_196, 3, x_195); -x_197 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__17; +x_197 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__18; lean_inc(x_123); x_198 = l_Lean_Syntax_node2(x_123, x_197, x_191, x_196); lean_inc(x_123); x_199 = l_Lean_Syntax_node1(x_123, x_173, x_198); -x_200 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__15; +x_200 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__16; lean_inc(x_123); x_201 = l_Lean_Syntax_node2(x_123, x_200, x_189, x_199); -x_202 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__27; +x_202 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__28; lean_inc(x_123); x_203 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_203, 0, x_123); lean_ctor_set(x_203, 1, x_202); -x_204 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__30; +x_204 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__31; lean_inc_n(x_175, 2); lean_inc(x_123); x_205 = l_Lean_Syntax_node2(x_123, x_204, x_175, x_175); -x_206 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__26; +x_206 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__27; lean_inc(x_175); lean_inc(x_123); x_207 = l_Lean_Syntax_node4(x_123, x_206, x_203, x_5, x_205, x_175); @@ -3633,13 +3642,13 @@ x_38 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__5; lean_inc_n(x_27, 4); lean_inc(x_18); x_39 = l_Lean_Syntax_node6(x_18, x_38, x_27, x_27, x_32, x_27, x_27, x_37); -x_40 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__10; +x_40 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__12; lean_inc(x_18); x_41 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_41, 0, x_18); lean_ctor_set(x_41, 1, x_40); x_42 = lean_mk_syntax_ident(x_2); -x_43 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__13; +x_43 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__14; lean_inc(x_27); lean_inc(x_18); x_44 = l_Lean_Syntax_node2(x_18, x_43, x_42, x_27); @@ -3649,39 +3658,39 @@ x_46 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_46, 0, x_18); lean_ctor_set(x_46, 1, x_25); lean_ctor_set(x_46, 2, x_45); -x_47 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__18; +x_47 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__19; lean_inc(x_18); x_48 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_48, 0, x_18); lean_ctor_set(x_48, 1, x_47); -x_49 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__20; +x_49 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__21; x_50 = l_Lean_addMacroScope(x_24, x_49, x_19); -x_51 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__19; -x_52 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__24; +x_51 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__20; +x_52 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__25; lean_inc(x_18); x_53 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_53, 0, x_18); lean_ctor_set(x_53, 1, x_51); lean_ctor_set(x_53, 2, x_50); lean_ctor_set(x_53, 3, x_52); -x_54 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__17; +x_54 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__18; lean_inc(x_18); x_55 = l_Lean_Syntax_node2(x_18, x_54, x_48, x_53); lean_inc(x_18); x_56 = l_Lean_Syntax_node1(x_18, x_25, x_55); -x_57 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__15; +x_57 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__16; lean_inc(x_18); x_58 = l_Lean_Syntax_node2(x_18, x_57, x_46, x_56); -x_59 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__27; +x_59 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__28; lean_inc(x_18); x_60 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_60, 0, x_18); lean_ctor_set(x_60, 1, x_59); -x_61 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__30; +x_61 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__31; lean_inc_n(x_27, 2); lean_inc(x_18); x_62 = l_Lean_Syntax_node2(x_18, x_61, x_27, x_27); -x_63 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__26; +x_63 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__27; lean_inc(x_27); lean_inc(x_18); x_64 = l_Lean_Syntax_node4(x_18, x_63, x_60, x_5, x_62, x_27); @@ -3736,13 +3745,13 @@ x_86 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__5; lean_inc_n(x_75, 4); lean_inc(x_18); x_87 = l_Lean_Syntax_node6(x_18, x_86, x_75, x_75, x_80, x_75, x_75, x_85); -x_88 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__10; +x_88 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__12; lean_inc(x_18); x_89 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_89, 0, x_18); lean_ctor_set(x_89, 1, x_88); x_90 = lean_mk_syntax_ident(x_2); -x_91 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__13; +x_91 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__14; lean_inc(x_75); lean_inc(x_18); x_92 = l_Lean_Syntax_node2(x_18, x_91, x_90, x_75); @@ -3752,39 +3761,39 @@ x_94 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_94, 0, x_18); lean_ctor_set(x_94, 1, x_73); lean_ctor_set(x_94, 2, x_93); -x_95 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__18; +x_95 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__19; lean_inc(x_18); x_96 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_96, 0, x_18); lean_ctor_set(x_96, 1, x_95); -x_97 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__20; +x_97 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__21; x_98 = l_Lean_addMacroScope(x_72, x_97, x_19); -x_99 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__19; -x_100 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__24; +x_99 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__20; +x_100 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__25; lean_inc(x_18); x_101 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_101, 0, x_18); lean_ctor_set(x_101, 1, x_99); lean_ctor_set(x_101, 2, x_98); lean_ctor_set(x_101, 3, x_100); -x_102 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__17; +x_102 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__18; lean_inc(x_18); x_103 = l_Lean_Syntax_node2(x_18, x_102, x_96, x_101); lean_inc(x_18); x_104 = l_Lean_Syntax_node1(x_18, x_73, x_103); -x_105 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__15; +x_105 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__16; lean_inc(x_18); x_106 = l_Lean_Syntax_node2(x_18, x_105, x_94, x_104); -x_107 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__27; +x_107 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__28; lean_inc(x_18); x_108 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_108, 0, x_18); lean_ctor_set(x_108, 1, x_107); -x_109 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__30; +x_109 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__31; lean_inc_n(x_75, 2); lean_inc(x_18); x_110 = l_Lean_Syntax_node2(x_18, x_109, x_75, x_75); -x_111 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__26; +x_111 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__27; lean_inc(x_75); lean_inc(x_18); x_112 = l_Lean_Syntax_node4(x_18, x_111, x_108, x_5, x_110, x_75); @@ -5804,6 +5813,8 @@ l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__29 = _init_l_Lean_E lean_mark_persistent(l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__29); l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__30 = _init_l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__30(); lean_mark_persistent(l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__30); +l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__31 = _init_l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__31(); +lean_mark_persistent(l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__31); l_Lean_Elab_Deriving_Ord_mkMutualBlock___closed__1 = _init_l_Lean_Elab_Deriving_Ord_mkMutualBlock___closed__1(); lean_mark_persistent(l_Lean_Elab_Deriving_Ord_mkMutualBlock___closed__1); l_Lean_Elab_Deriving_Ord_mkMutualBlock___closed__2 = _init_l_Lean_Elab_Deriving_Ord_mkMutualBlock___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/Deriving/Repr.c b/stage0/stdlib/Lean/Elab/Deriving/Repr.c index c32b47704112..b08c7efeae89 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/Repr.c +++ b/stage0/stdlib/Lean/Elab/Deriving/Repr.c @@ -258,6 +258,7 @@ static lean_object* l_Lean_Elab_Deriving_Repr_mkMutualBlock___closed__3; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__26; static lean_object* l___private_Lean_Elab_Deriving_Repr_0__Lean_Elab_Deriving_Repr_mkReprInstanceCmd___closed__6; static lean_object* l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__20; +static lean_object* l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__31; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__8; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Deriving_Repr_mkReprInstanceHandler___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -5146,7 +5147,7 @@ static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___ _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("def", 3); +x_1 = lean_mk_string_from_bytes("definition", 10); return x_1; } } @@ -5166,23 +5167,31 @@ static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___ _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("declId", 6); +x_1 = lean_mk_string_from_bytes("def", 3); return x_1; } } static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__11() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("declId", 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkAuxFunction___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_Lean_Elab_Deriving_Repr_mkReprHeader___closed__3; x_2 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__4; x_3 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__1; -x_4 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__10; +x_4 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___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_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__12() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__13() { _start: { lean_object* x_1; @@ -5190,19 +5199,19 @@ x_1 = lean_mk_string_from_bytes("optDeclSig", 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__13() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___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_Elab_Deriving_Repr_mkReprHeader___closed__3; x_2 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__4; x_3 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__1; -x_4 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__12; +x_4 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___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_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__14() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__15() { _start: { lean_object* x_1; @@ -5210,19 +5219,19 @@ x_1 = lean_mk_string_from_bytes("typeSpec", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__15() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___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_Elab_Deriving_Repr_mkReprHeader___closed__3; x_2 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__4; x_3 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__5; -x_4 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__14; +x_4 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___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_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__16() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__17() { _start: { lean_object* x_1; lean_object* x_2; @@ -5231,7 +5240,7 @@ x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__17() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -5241,7 +5250,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__18() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -5251,53 +5260,53 @@ x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__19() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkAuxFunction___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_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__18; +x_2 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___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_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__20() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__21() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__18; +x_1 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___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_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__21() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__20; +x_2 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__21; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__22() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__19; -x_2 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__21; +x_1 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__20; +x_2 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___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_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__23() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__24() { _start: { lean_object* x_1; @@ -5305,19 +5314,19 @@ x_1 = lean_mk_string_from_bytes("declValSimple", 13); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__24() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__25() { _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_Repr_mkReprHeader___closed__3; x_2 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__4; x_3 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__1; -x_4 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__23; +x_4 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__24; 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_Repr_mkAuxFunction___lambda__1___closed__25() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__26() { _start: { lean_object* x_1; @@ -5325,7 +5334,7 @@ x_1 = lean_mk_string_from_bytes(":=", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__26() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__27() { _start: { lean_object* x_1; @@ -5333,7 +5342,7 @@ x_1 = lean_mk_string_from_bytes("Termination", 11); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__27() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__28() { _start: { lean_object* x_1; @@ -5341,19 +5350,19 @@ x_1 = lean_mk_string_from_bytes("suffix", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__28() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__29() { _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_Repr_mkReprHeader___closed__3; x_2 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__4; -x_3 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__26; -x_4 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__27; +x_3 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__27; +x_4 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__28; 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_Repr_mkAuxFunction___lambda__1___closed__29() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__30() { _start: { lean_object* x_1; @@ -5361,14 +5370,14 @@ x_1 = lean_mk_string_from_bytes("partial", 7); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__30() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__31() { _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_Repr_mkReprHeader___closed__3; x_2 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__4; x_3 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__1; -x_4 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__29; +x_4 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__30; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } @@ -5422,13 +5431,13 @@ x_32 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__5; lean_inc_n(x_26, 5); lean_inc(x_17); x_33 = l_Lean_Syntax_node6(x_17, x_32, x_26, x_26, x_31, x_26, x_26, x_26); -x_34 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__8; +x_34 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__10; lean_inc(x_17); x_35 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_35, 0, x_17); lean_ctor_set(x_35, 1, x_34); x_36 = lean_mk_syntax_ident(x_3); -x_37 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__11; +x_37 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__12; lean_inc(x_26); lean_inc(x_17); x_38 = l_Lean_Syntax_node2(x_17, x_37, x_36, x_26); @@ -5443,34 +5452,34 @@ lean_inc(x_17); x_42 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_42, 0, x_17); lean_ctor_set(x_42, 1, x_41); -x_43 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__17; +x_43 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__18; x_44 = l_Lean_addMacroScope(x_23, x_43, x_18); -x_45 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__16; -x_46 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__22; +x_45 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__17; +x_46 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__23; lean_inc(x_17); x_47 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_47, 0, x_17); lean_ctor_set(x_47, 1, x_45); lean_ctor_set(x_47, 2, x_44); lean_ctor_set(x_47, 3, x_46); -x_48 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__15; +x_48 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__16; lean_inc(x_17); x_49 = l_Lean_Syntax_node2(x_17, x_48, x_42, x_47); lean_inc(x_17); x_50 = l_Lean_Syntax_node1(x_17, x_24, x_49); -x_51 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__13; +x_51 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__14; lean_inc(x_17); x_52 = l_Lean_Syntax_node2(x_17, x_51, x_40, x_50); -x_53 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__25; +x_53 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__26; lean_inc(x_17); x_54 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_54, 0, x_17); lean_ctor_set(x_54, 1, x_53); -x_55 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__28; +x_55 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__29; lean_inc_n(x_26, 2); lean_inc(x_17); x_56 = l_Lean_Syntax_node2(x_17, x_55, x_26, x_26); -x_57 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__24; +x_57 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__25; lean_inc(x_26); lean_inc(x_17); x_58 = l_Lean_Syntax_node4(x_17, x_57, x_54, x_4, x_56, x_26); @@ -5515,13 +5524,13 @@ x_75 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__5; lean_inc_n(x_69, 5); lean_inc(x_17); x_76 = l_Lean_Syntax_node6(x_17, x_75, x_69, x_69, x_74, x_69, x_69, x_69); -x_77 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__8; +x_77 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__10; lean_inc(x_17); x_78 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_78, 0, x_17); lean_ctor_set(x_78, 1, x_77); x_79 = lean_mk_syntax_ident(x_3); -x_80 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__11; +x_80 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__12; lean_inc(x_69); lean_inc(x_17); x_81 = l_Lean_Syntax_node2(x_17, x_80, x_79, x_69); @@ -5536,34 +5545,34 @@ lean_inc(x_17); x_85 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_85, 0, x_17); lean_ctor_set(x_85, 1, x_84); -x_86 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__17; +x_86 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__18; x_87 = l_Lean_addMacroScope(x_66, x_86, x_18); -x_88 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__16; -x_89 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__22; +x_88 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__17; +x_89 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__23; lean_inc(x_17); x_90 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_90, 0, x_17); lean_ctor_set(x_90, 1, x_88); lean_ctor_set(x_90, 2, x_87); lean_ctor_set(x_90, 3, x_89); -x_91 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__15; +x_91 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__16; lean_inc(x_17); x_92 = l_Lean_Syntax_node2(x_17, x_91, x_85, x_90); lean_inc(x_17); x_93 = l_Lean_Syntax_node1(x_17, x_67, x_92); -x_94 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__13; +x_94 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__14; lean_inc(x_17); x_95 = l_Lean_Syntax_node2(x_17, x_94, x_83, x_93); -x_96 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__25; +x_96 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__26; lean_inc(x_17); x_97 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_97, 0, x_17); lean_ctor_set(x_97, 1, x_96); -x_98 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__28; +x_98 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__29; lean_inc_n(x_69, 2); lean_inc(x_17); x_99 = l_Lean_Syntax_node2(x_17, x_98, x_69, x_69); -x_100 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__24; +x_100 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__25; lean_inc(x_69); lean_inc(x_17); x_101 = l_Lean_Syntax_node4(x_17, x_100, x_97, x_4, x_99, x_69); @@ -5618,12 +5627,12 @@ lean_inc(x_110); x_123 = l_Lean_Syntax_node1(x_110, x_122, x_121); lean_inc(x_110); x_124 = l_Lean_Syntax_node1(x_110, x_117, x_123); -x_125 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__29; +x_125 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__30; lean_inc(x_110); x_126 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_126, 0, x_110); lean_ctor_set(x_126, 1, x_125); -x_127 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__30; +x_127 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__31; lean_inc(x_110); x_128 = l_Lean_Syntax_node1(x_110, x_127, x_126); lean_inc(x_110); @@ -5632,13 +5641,13 @@ x_130 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__5; lean_inc_n(x_119, 4); lean_inc(x_110); x_131 = l_Lean_Syntax_node6(x_110, x_130, x_119, x_119, x_124, x_119, x_119, x_129); -x_132 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__8; +x_132 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__10; lean_inc(x_110); x_133 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_133, 0, x_110); lean_ctor_set(x_133, 1, x_132); x_134 = lean_mk_syntax_ident(x_3); -x_135 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__11; +x_135 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__12; lean_inc(x_119); lean_inc(x_110); x_136 = l_Lean_Syntax_node2(x_110, x_135, x_134, x_119); @@ -5653,34 +5662,34 @@ lean_inc(x_110); x_140 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_140, 0, x_110); lean_ctor_set(x_140, 1, x_139); -x_141 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__17; +x_141 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__18; x_142 = l_Lean_addMacroScope(x_116, x_141, x_111); -x_143 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__16; -x_144 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__22; +x_143 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__17; +x_144 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__23; lean_inc(x_110); x_145 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_145, 0, x_110); lean_ctor_set(x_145, 1, x_143); lean_ctor_set(x_145, 2, x_142); lean_ctor_set(x_145, 3, x_144); -x_146 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__15; +x_146 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__16; lean_inc(x_110); x_147 = l_Lean_Syntax_node2(x_110, x_146, x_140, x_145); lean_inc(x_110); x_148 = l_Lean_Syntax_node1(x_110, x_117, x_147); -x_149 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__13; +x_149 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__14; lean_inc(x_110); x_150 = l_Lean_Syntax_node2(x_110, x_149, x_138, x_148); -x_151 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__25; +x_151 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__26; lean_inc(x_110); x_152 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_152, 0, x_110); lean_ctor_set(x_152, 1, x_151); -x_153 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__28; +x_153 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__29; lean_inc_n(x_119, 2); lean_inc(x_110); x_154 = l_Lean_Syntax_node2(x_110, x_153, x_119, x_119); -x_155 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__24; +x_155 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__25; lean_inc(x_119); lean_inc(x_110); x_156 = l_Lean_Syntax_node4(x_110, x_155, x_152, x_4, x_154, x_119); @@ -5721,12 +5730,12 @@ lean_inc(x_110); x_171 = l_Lean_Syntax_node1(x_110, x_170, x_169); lean_inc(x_110); x_172 = l_Lean_Syntax_node1(x_110, x_165, x_171); -x_173 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__29; +x_173 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__30; lean_inc(x_110); x_174 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_174, 0, x_110); lean_ctor_set(x_174, 1, x_173); -x_175 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__30; +x_175 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__31; lean_inc(x_110); x_176 = l_Lean_Syntax_node1(x_110, x_175, x_174); lean_inc(x_110); @@ -5735,13 +5744,13 @@ x_178 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__5; lean_inc_n(x_167, 4); lean_inc(x_110); x_179 = l_Lean_Syntax_node6(x_110, x_178, x_167, x_167, x_172, x_167, x_167, x_177); -x_180 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__8; +x_180 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__10; lean_inc(x_110); x_181 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_181, 0, x_110); lean_ctor_set(x_181, 1, x_180); x_182 = lean_mk_syntax_ident(x_3); -x_183 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__11; +x_183 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__12; lean_inc(x_167); lean_inc(x_110); x_184 = l_Lean_Syntax_node2(x_110, x_183, x_182, x_167); @@ -5756,34 +5765,34 @@ lean_inc(x_110); x_188 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_188, 0, x_110); lean_ctor_set(x_188, 1, x_187); -x_189 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__17; +x_189 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__18; x_190 = l_Lean_addMacroScope(x_164, x_189, x_111); -x_191 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__16; -x_192 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__22; +x_191 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__17; +x_192 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__23; lean_inc(x_110); x_193 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_193, 0, x_110); lean_ctor_set(x_193, 1, x_191); lean_ctor_set(x_193, 2, x_190); lean_ctor_set(x_193, 3, x_192); -x_194 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__15; +x_194 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__16; lean_inc(x_110); x_195 = l_Lean_Syntax_node2(x_110, x_194, x_188, x_193); lean_inc(x_110); x_196 = l_Lean_Syntax_node1(x_110, x_165, x_195); -x_197 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__13; +x_197 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__14; lean_inc(x_110); x_198 = l_Lean_Syntax_node2(x_110, x_197, x_186, x_196); -x_199 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__25; +x_199 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__26; lean_inc(x_110); x_200 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_200, 0, x_110); lean_ctor_set(x_200, 1, x_199); -x_201 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__28; +x_201 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__29; lean_inc_n(x_167, 2); lean_inc(x_110); x_202 = l_Lean_Syntax_node2(x_110, x_201, x_167, x_167); -x_203 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__24; +x_203 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__25; lean_inc(x_167); lean_inc(x_110); x_204 = l_Lean_Syntax_node4(x_110, x_203, x_200, x_4, x_202, x_167); @@ -7855,6 +7864,8 @@ l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__29 = _init_l_Lean_ lean_mark_persistent(l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__29); l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__30 = _init_l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__30(); lean_mark_persistent(l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__30); +l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__31 = _init_l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__31(); +lean_mark_persistent(l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__31); l_Lean_Elab_Deriving_Repr_mkMutualBlock___closed__1 = _init_l_Lean_Elab_Deriving_Repr_mkMutualBlock___closed__1(); lean_mark_persistent(l_Lean_Elab_Deriving_Repr_mkMutualBlock___closed__1); l_Lean_Elab_Deriving_Repr_mkMutualBlock___closed__2 = _init_l_Lean_Elab_Deriving_Repr_mkMutualBlock___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/Deriving/TypeName.c b/stage0/stdlib/Lean/Elab/Deriving/TypeName.c index a3b92895775e..765dc3fdd3e7 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/TypeName.c +++ b/stage0/stdlib/Lean/Elab/Deriving/TypeName.c @@ -110,6 +110,7 @@ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_T static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__76; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__16; lean_object* l_Lean_Syntax_node1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__80; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__79; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__24; @@ -259,7 +260,7 @@ static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deri _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("def", 3); +x_1 = lean_mk_string_from_bytes("definition", 10); return x_1; } } @@ -279,23 +280,31 @@ static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deri _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("declId", 6); +x_1 = lean_mk_string_from_bytes("def", 3); return x_1; } } static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__16() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("declId", 6); +return x_1; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___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_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__3; x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__4; x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__5; -x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__15; +x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__16; 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___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__17() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__18() { _start: { lean_object* x_1; @@ -303,26 +312,26 @@ x_1 = lean_mk_string_from_bytes("instImpl", 8); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__18() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__19() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__17; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__18; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__19() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__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_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__17; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__18; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__20() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__21() { _start: { lean_object* x_1; @@ -330,19 +339,19 @@ x_1 = lean_mk_string_from_bytes("optDeclSig", 10); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__21() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__22() { _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_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__3; x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__4; x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__5; -x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__20; +x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__21; 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___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__22() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__23() { _start: { lean_object* x_1; @@ -350,7 +359,7 @@ x_1 = lean_mk_string_from_bytes("Term", 4); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__23() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__24() { _start: { lean_object* x_1; @@ -358,19 +367,19 @@ x_1 = lean_mk_string_from_bytes("typeSpec", 8); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__24() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__25() { _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_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__3; x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__4; -x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__22; -x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__23; +x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__23; +x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__24; 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___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__25() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__26() { _start: { lean_object* x_1; @@ -378,7 +387,7 @@ x_1 = lean_mk_string_from_bytes(":", 1); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__26() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__27() { _start: { lean_object* x_1; @@ -386,19 +395,19 @@ x_1 = lean_mk_string_from_bytes("app", 3); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__27() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___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_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__3; x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__4; -x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__22; -x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__26; +x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__23; +x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__27; 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___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__28() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__29() { _start: { lean_object* x_1; @@ -406,72 +415,72 @@ x_1 = lean_mk_string_from_bytes("TypeName", 8); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__29() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__30() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__28; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__29; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__30() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__31() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__28; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__29; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__31() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__32() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__30; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__31; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__32() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__33() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__30; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__31; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__33() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__34() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__32; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__33; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__34() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__35() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__31; -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__33; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__32; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___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_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__35() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__36() { _start: { lean_object* x_1; @@ -479,19 +488,19 @@ x_1 = lean_mk_string_from_bytes("explicit", 8); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__36() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__37() { _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_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__3; x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__4; -x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__22; -x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__35; +x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__23; +x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__36; 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___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__37() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__38() { _start: { lean_object* x_1; @@ -499,7 +508,7 @@ x_1 = lean_mk_string_from_bytes("@", 1); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__38() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__39() { _start: { lean_object* x_1; @@ -507,19 +516,19 @@ x_1 = lean_mk_string_from_bytes("declValSimple", 13); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__39() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__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_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__3; x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__4; x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__5; -x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__38; +x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__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_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__40() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__41() { _start: { lean_object* x_1; @@ -527,7 +536,7 @@ x_1 = lean_mk_string_from_bytes(":=", 2); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__41() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__42() { _start: { lean_object* x_1; @@ -535,19 +544,19 @@ x_1 = lean_mk_string_from_bytes("dotIdent", 8); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__42() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__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_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__3; x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__4; -x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__22; -x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__41; +x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__23; +x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__42; 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___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__43() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__44() { _start: { lean_object* x_1; @@ -555,7 +564,7 @@ x_1 = lean_mk_string_from_bytes(".", 1); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__44() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__45() { _start: { lean_object* x_1; @@ -563,26 +572,26 @@ x_1 = lean_mk_string_from_bytes("mk", 2); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__45() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__46() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__44; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__45; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__46() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__47() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__44; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__45; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__47() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__48() { _start: { lean_object* x_1; @@ -590,19 +599,19 @@ x_1 = lean_mk_string_from_bytes("hole", 4); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__48() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__49() { _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_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__3; x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__4; -x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__22; -x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__47; +x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__23; +x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__48; 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___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__49() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__50() { _start: { lean_object* x_1; @@ -610,7 +619,7 @@ x_1 = lean_mk_string_from_bytes("_", 1); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__50() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__51() { _start: { lean_object* x_1; @@ -618,7 +627,7 @@ x_1 = lean_mk_string_from_bytes("Termination", 11); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__51() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__52() { _start: { lean_object* x_1; @@ -626,19 +635,19 @@ x_1 = lean_mk_string_from_bytes("suffix", 6); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__52() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___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_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__3; x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__4; -x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__50; -x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__51; +x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__51; +x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__52; 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___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__53() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__54() { _start: { lean_object* x_1; @@ -646,19 +655,19 @@ x_1 = lean_mk_string_from_bytes("attributes", 10); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__54() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__55() { _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_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__3; x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__4; -x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__22; -x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__53; +x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__23; +x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__54; 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___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__55() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__56() { _start: { lean_object* x_1; @@ -666,7 +675,7 @@ x_1 = lean_mk_string_from_bytes("@[", 2); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__56() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__57() { _start: { lean_object* x_1; @@ -674,19 +683,19 @@ x_1 = lean_mk_string_from_bytes("attrInstance", 12); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__57() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__58() { _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_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__3; x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__4; -x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__22; -x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__56; +x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__23; +x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__57; 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___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__58() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__59() { _start: { lean_object* x_1; @@ -694,19 +703,19 @@ x_1 = lean_mk_string_from_bytes("attrKind", 8); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__59() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__60() { _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_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__3; x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__4; -x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__22; -x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__58; +x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__23; +x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__59; 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___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__60() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__61() { _start: { lean_object* x_1; @@ -714,7 +723,7 @@ x_1 = lean_mk_string_from_bytes("Attr", 4); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__61() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__62() { _start: { lean_object* x_1; @@ -722,19 +731,19 @@ x_1 = lean_mk_string_from_bytes("simple", 6); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__62() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___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_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__3; x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__4; -x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__60; -x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__61; +x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__61; +x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__62; 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___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__63() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__64() { _start: { lean_object* x_1; @@ -742,26 +751,26 @@ x_1 = lean_mk_string_from_bytes("implemented_by", 14); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__64() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__65() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__63; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__64; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__65() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__66() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__63; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__64; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__66() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__67() { _start: { lean_object* x_1; @@ -769,7 +778,7 @@ x_1 = lean_mk_string_from_bytes("]", 1); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__67() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__68() { _start: { lean_object* x_1; @@ -777,19 +786,19 @@ x_1 = lean_mk_string_from_bytes("opaque", 6); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__68() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__69() { _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_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__3; x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__4; x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__5; -x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__67; +x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__68; 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___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__69() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__70() { _start: { lean_object* x_1; @@ -797,26 +806,26 @@ x_1 = lean_mk_string_from_bytes("inst", 4); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__70() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__71() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__69; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__70; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__71() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__72() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__69; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__70; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__72() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__73() { _start: { lean_object* x_1; @@ -824,19 +833,19 @@ x_1 = lean_mk_string_from_bytes("declSig", 7); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__73() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__74() { _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_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__3; x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__4; x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__5; -x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__72; +x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__73; 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___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__74() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__75() { _start: { lean_object* x_1; @@ -844,19 +853,19 @@ x_1 = lean_mk_string_from_bytes("instance", 8); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__75() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__76() { _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_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__3; x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__4; x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__5; -x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__74; +x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__75; 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___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__76() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__77() { _start: { lean_object* x_1; @@ -864,19 +873,19 @@ x_1 = lean_mk_string_from_bytes("quotedName", 10); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__77() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__78() { _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_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__3; x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__4; -x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__22; -x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__76; +x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__23; +x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__77; 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___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__78() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__79() { _start: { lean_object* x_1; @@ -884,7 +893,7 @@ x_1 = lean_mk_string_from_bytes("`", 1); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__79() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__80() { _start: { lean_object* x_1; lean_object* x_2; @@ -938,46 +947,46 @@ x_24 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Le lean_inc_n(x_18, 5); lean_inc(x_9); x_25 = l_Lean_Syntax_node6(x_9, x_24, x_18, x_18, x_18, x_18, x_23, x_18); -x_26 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__13; +x_26 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__15; lean_inc(x_9); x_27 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_27, 0, x_9); lean_ctor_set(x_27, 1, x_26); -x_28 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__19; +x_28 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__20; lean_inc(x_11); lean_inc(x_15); x_29 = l_Lean_addMacroScope(x_15, x_28, x_11); x_30 = lean_box(0); -x_31 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__18; +x_31 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__19; lean_inc(x_9); x_32 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_32, 0, x_9); lean_ctor_set(x_32, 1, x_31); lean_ctor_set(x_32, 2, x_29); lean_ctor_set(x_32, 3, x_30); -x_33 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__16; +x_33 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__17; lean_inc(x_18); lean_inc(x_32); lean_inc(x_9); x_34 = l_Lean_Syntax_node2(x_9, x_33, x_32, x_18); -x_35 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__25; +x_35 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__26; lean_inc(x_9); x_36 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_36, 0, x_9); lean_ctor_set(x_36, 1, x_35); -x_37 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__30; +x_37 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__31; lean_inc(x_11); lean_inc(x_15); x_38 = l_Lean_addMacroScope(x_15, x_37, x_11); -x_39 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__29; -x_40 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__34; +x_39 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__30; +x_40 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__35; lean_inc(x_9); x_41 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_41, 0, x_9); lean_ctor_set(x_41, 1, x_39); lean_ctor_set(x_41, 2, x_38); lean_ctor_set(x_41, 3, x_40); -x_42 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__37; +x_42 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__38; lean_inc(x_9); x_43 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_43, 0, x_9); @@ -985,76 +994,76 @@ lean_ctor_set(x_43, 1, x_42); x_44 = lean_box(0); lean_inc(x_1); x_45 = l_Lean_mkCIdentFrom(x_44, x_1, x_8); -x_46 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__36; +x_46 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__37; lean_inc(x_9); x_47 = l_Lean_Syntax_node2(x_9, x_46, x_43, x_45); lean_inc(x_9); x_48 = l_Lean_Syntax_node1(x_9, x_16, x_47); -x_49 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__27; +x_49 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__28; lean_inc(x_9); x_50 = l_Lean_Syntax_node2(x_9, x_49, x_41, x_48); -x_51 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__24; +x_51 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__25; lean_inc(x_9); x_52 = l_Lean_Syntax_node2(x_9, x_51, x_36, x_50); lean_inc(x_52); lean_inc(x_9); x_53 = l_Lean_Syntax_node1(x_9, x_16, x_52); -x_54 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__21; +x_54 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__22; lean_inc(x_18); lean_inc(x_9); x_55 = l_Lean_Syntax_node2(x_9, x_54, x_18, x_53); -x_56 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__40; +x_56 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__41; lean_inc(x_9); x_57 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_57, 0, x_9); lean_ctor_set(x_57, 1, x_56); -x_58 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__43; +x_58 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__44; lean_inc(x_9); x_59 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_59, 0, x_9); lean_ctor_set(x_59, 1, x_58); -x_60 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__46; +x_60 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__47; lean_inc(x_11); lean_inc(x_15); x_61 = l_Lean_addMacroScope(x_15, x_60, x_11); -x_62 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__45; +x_62 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__46; lean_inc(x_9); x_63 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_63, 0, x_9); lean_ctor_set(x_63, 1, x_62); lean_ctor_set(x_63, 2, x_61); lean_ctor_set(x_63, 3, x_30); -x_64 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__42; +x_64 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__43; lean_inc(x_9); x_65 = l_Lean_Syntax_node2(x_9, x_64, x_59, x_63); -x_66 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__49; +x_66 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__50; lean_inc(x_9); x_67 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_67, 0, x_9); lean_ctor_set(x_67, 1, x_66); -x_68 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__48; +x_68 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__49; lean_inc(x_9); x_69 = l_Lean_Syntax_node1(x_9, x_68, x_67); lean_inc(x_1); x_70 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_30, x_1); -x_71 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__52; +x_71 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__53; lean_inc_n(x_18, 2); lean_inc(x_9); x_72 = l_Lean_Syntax_node2(x_9, x_71, x_18, x_18); -x_73 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__55; +x_73 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__56; lean_inc(x_9); x_74 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_74, 0, x_9); lean_ctor_set(x_74, 1, x_73); -x_75 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__59; +x_75 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__60; lean_inc(x_18); lean_inc(x_9); x_76 = l_Lean_Syntax_node1(x_9, x_75, x_18); -x_77 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__65; +x_77 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__66; lean_inc(x_11); lean_inc(x_15); x_78 = l_Lean_addMacroScope(x_15, x_77, x_11); -x_79 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__64; +x_79 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__65; lean_inc(x_9); x_80 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_80, 0, x_9); @@ -1063,21 +1072,21 @@ lean_ctor_set(x_80, 2, x_78); lean_ctor_set(x_80, 3, x_30); lean_inc(x_9); x_81 = l_Lean_Syntax_node1(x_9, x_16, x_32); -x_82 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__62; +x_82 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__63; lean_inc(x_9); x_83 = l_Lean_Syntax_node2(x_9, x_82, x_80, x_81); -x_84 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__57; +x_84 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__58; lean_inc(x_76); lean_inc(x_9); x_85 = l_Lean_Syntax_node2(x_9, x_84, x_76, x_83); lean_inc(x_9); x_86 = l_Lean_Syntax_node1(x_9, x_16, x_85); -x_87 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__66; +x_87 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__67; lean_inc(x_9); x_88 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_88, 0, x_9); lean_ctor_set(x_88, 1, x_87); -x_89 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__54; +x_89 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__55; lean_inc(x_9); x_90 = l_Lean_Syntax_node3(x_9, x_89, x_74, x_86, x_88); lean_inc(x_9); @@ -1085,14 +1094,14 @@ x_91 = l_Lean_Syntax_node1(x_9, x_16, x_90); lean_inc_n(x_18, 5); lean_inc(x_9); x_92 = l_Lean_Syntax_node6(x_9, x_24, x_18, x_91, x_18, x_18, x_18, x_18); -x_93 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__67; +x_93 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__68; lean_inc(x_9); x_94 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_94, 0, x_9); lean_ctor_set(x_94, 1, x_93); -x_95 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__71; +x_95 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__72; x_96 = l_Lean_addMacroScope(x_15, x_95, x_11); -x_97 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__70; +x_97 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__71; lean_inc(x_9); x_98 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_98, 0, x_9); @@ -1103,11 +1112,11 @@ lean_inc(x_18); lean_inc(x_98); lean_inc(x_9); x_99 = l_Lean_Syntax_node2(x_9, x_33, x_98, x_18); -x_100 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__73; +x_100 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__74; lean_inc(x_18); lean_inc(x_9); x_101 = l_Lean_Syntax_node2(x_9, x_100, x_18, x_52); -x_102 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__68; +x_102 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__69; lean_inc(x_18); lean_inc(x_101); lean_inc(x_9); @@ -1118,18 +1127,18 @@ x_105 = l_Lean_Syntax_node2(x_9, x_104, x_92, x_103); lean_inc_n(x_18, 6); lean_inc(x_9); x_106 = l_Lean_Syntax_node6(x_9, x_24, x_18, x_18, x_18, x_18, x_18, x_18); -x_107 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__74; +x_107 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__75; lean_inc(x_9); x_108 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_108, 0, x_9); lean_ctor_set(x_108, 1, x_107); -x_109 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__39; +x_109 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__40; lean_inc(x_18); lean_inc(x_72); lean_inc(x_57); lean_inc(x_9); x_110 = l_Lean_Syntax_node4(x_9, x_109, x_57, x_98, x_72, x_18); -x_111 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__75; +x_111 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__76; lean_inc_n(x_18, 2); lean_inc(x_9); x_112 = l_Lean_Syntax_node6(x_9, x_111, x_76, x_108, x_18, x_18, x_101, x_110); @@ -1163,14 +1172,14 @@ x_122 = lean_ctor_get(x_70, 0); lean_inc(x_122); lean_dec(x_70); x_123 = l_String_intercalate(x_58, x_122); -x_124 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__78; +x_124 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__79; x_125 = lean_string_append(x_124, x_123); lean_dec(x_123); x_126 = lean_box(2); x_127 = l_Lean_Syntax_mkNameLit(x_125, x_126); -x_128 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__79; +x_128 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__80; x_129 = lean_array_push(x_128, x_127); -x_130 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__77; +x_130 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__78; x_131 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_131, 0, x_126); lean_ctor_set(x_131, 1, x_130); @@ -1221,46 +1230,46 @@ x_149 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__L lean_inc_n(x_143, 5); lean_inc(x_9); x_150 = l_Lean_Syntax_node6(x_9, x_149, x_143, x_143, x_143, x_143, x_148, x_143); -x_151 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__13; +x_151 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__15; lean_inc(x_9); x_152 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_152, 0, x_9); lean_ctor_set(x_152, 1, x_151); -x_153 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__19; +x_153 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__20; lean_inc(x_11); lean_inc(x_139); x_154 = l_Lean_addMacroScope(x_139, x_153, x_11); x_155 = lean_box(0); -x_156 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__18; +x_156 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__19; lean_inc(x_9); x_157 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_157, 0, x_9); lean_ctor_set(x_157, 1, x_156); lean_ctor_set(x_157, 2, x_154); lean_ctor_set(x_157, 3, x_155); -x_158 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__16; +x_158 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__17; lean_inc(x_143); lean_inc(x_157); lean_inc(x_9); x_159 = l_Lean_Syntax_node2(x_9, x_158, x_157, x_143); -x_160 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__25; +x_160 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__26; lean_inc(x_9); x_161 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_161, 0, x_9); lean_ctor_set(x_161, 1, x_160); -x_162 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__30; +x_162 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__31; lean_inc(x_11); lean_inc(x_139); x_163 = l_Lean_addMacroScope(x_139, x_162, x_11); -x_164 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__29; -x_165 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__34; +x_164 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__30; +x_165 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__35; lean_inc(x_9); x_166 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_166, 0, x_9); lean_ctor_set(x_166, 1, x_164); lean_ctor_set(x_166, 2, x_163); lean_ctor_set(x_166, 3, x_165); -x_167 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__37; +x_167 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__38; lean_inc(x_9); x_168 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_168, 0, x_9); @@ -1268,76 +1277,76 @@ lean_ctor_set(x_168, 1, x_167); x_169 = lean_box(0); lean_inc(x_1); x_170 = l_Lean_mkCIdentFrom(x_169, x_1, x_8); -x_171 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__36; +x_171 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__37; lean_inc(x_9); x_172 = l_Lean_Syntax_node2(x_9, x_171, x_168, x_170); lean_inc(x_9); x_173 = l_Lean_Syntax_node1(x_9, x_141, x_172); -x_174 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__27; +x_174 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__28; lean_inc(x_9); x_175 = l_Lean_Syntax_node2(x_9, x_174, x_166, x_173); -x_176 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__24; +x_176 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__25; lean_inc(x_9); x_177 = l_Lean_Syntax_node2(x_9, x_176, x_161, x_175); lean_inc(x_177); lean_inc(x_9); x_178 = l_Lean_Syntax_node1(x_9, x_141, x_177); -x_179 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__21; +x_179 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__22; lean_inc(x_143); lean_inc(x_9); x_180 = l_Lean_Syntax_node2(x_9, x_179, x_143, x_178); -x_181 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__40; +x_181 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__41; lean_inc(x_9); x_182 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_182, 0, x_9); lean_ctor_set(x_182, 1, x_181); -x_183 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__43; +x_183 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__44; lean_inc(x_9); x_184 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_184, 0, x_9); lean_ctor_set(x_184, 1, x_183); -x_185 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__46; +x_185 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__47; lean_inc(x_11); lean_inc(x_139); x_186 = l_Lean_addMacroScope(x_139, x_185, x_11); -x_187 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__45; +x_187 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__46; lean_inc(x_9); x_188 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_188, 0, x_9); lean_ctor_set(x_188, 1, x_187); lean_ctor_set(x_188, 2, x_186); lean_ctor_set(x_188, 3, x_155); -x_189 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__42; +x_189 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__43; lean_inc(x_9); x_190 = l_Lean_Syntax_node2(x_9, x_189, x_184, x_188); -x_191 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__49; +x_191 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__50; lean_inc(x_9); x_192 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_192, 0, x_9); lean_ctor_set(x_192, 1, x_191); -x_193 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__48; +x_193 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__49; lean_inc(x_9); x_194 = l_Lean_Syntax_node1(x_9, x_193, x_192); lean_inc(x_1); x_195 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_155, x_1); -x_196 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__52; +x_196 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__53; lean_inc_n(x_143, 2); lean_inc(x_9); x_197 = l_Lean_Syntax_node2(x_9, x_196, x_143, x_143); -x_198 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__55; +x_198 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__56; lean_inc(x_9); x_199 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_199, 0, x_9); lean_ctor_set(x_199, 1, x_198); -x_200 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__59; +x_200 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__60; lean_inc(x_143); lean_inc(x_9); x_201 = l_Lean_Syntax_node1(x_9, x_200, x_143); -x_202 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__65; +x_202 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__66; lean_inc(x_11); lean_inc(x_139); x_203 = l_Lean_addMacroScope(x_139, x_202, x_11); -x_204 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__64; +x_204 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__65; lean_inc(x_9); x_205 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_205, 0, x_9); @@ -1346,21 +1355,21 @@ lean_ctor_set(x_205, 2, x_203); lean_ctor_set(x_205, 3, x_155); lean_inc(x_9); x_206 = l_Lean_Syntax_node1(x_9, x_141, x_157); -x_207 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__62; +x_207 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__63; lean_inc(x_9); x_208 = l_Lean_Syntax_node2(x_9, x_207, x_205, x_206); -x_209 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__57; +x_209 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__58; lean_inc(x_201); lean_inc(x_9); x_210 = l_Lean_Syntax_node2(x_9, x_209, x_201, x_208); lean_inc(x_9); x_211 = l_Lean_Syntax_node1(x_9, x_141, x_210); -x_212 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__66; +x_212 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__67; lean_inc(x_9); x_213 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_213, 0, x_9); lean_ctor_set(x_213, 1, x_212); -x_214 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__54; +x_214 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__55; lean_inc(x_9); x_215 = l_Lean_Syntax_node3(x_9, x_214, x_199, x_211, x_213); lean_inc(x_9); @@ -1368,14 +1377,14 @@ x_216 = l_Lean_Syntax_node1(x_9, x_141, x_215); lean_inc_n(x_143, 5); lean_inc(x_9); x_217 = l_Lean_Syntax_node6(x_9, x_149, x_143, x_216, x_143, x_143, x_143, x_143); -x_218 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__67; +x_218 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__68; lean_inc(x_9); x_219 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_219, 0, x_9); lean_ctor_set(x_219, 1, x_218); -x_220 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__71; +x_220 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__72; x_221 = l_Lean_addMacroScope(x_139, x_220, x_11); -x_222 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__70; +x_222 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__71; lean_inc(x_9); x_223 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_223, 0, x_9); @@ -1386,11 +1395,11 @@ lean_inc(x_143); lean_inc(x_223); lean_inc(x_9); x_224 = l_Lean_Syntax_node2(x_9, x_158, x_223, x_143); -x_225 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__73; +x_225 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__74; lean_inc(x_143); lean_inc(x_9); x_226 = l_Lean_Syntax_node2(x_9, x_225, x_143, x_177); -x_227 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__68; +x_227 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__69; lean_inc(x_143); lean_inc(x_226); lean_inc(x_9); @@ -1401,18 +1410,18 @@ x_230 = l_Lean_Syntax_node2(x_9, x_229, x_217, x_228); lean_inc_n(x_143, 6); lean_inc(x_9); x_231 = l_Lean_Syntax_node6(x_9, x_149, x_143, x_143, x_143, x_143, x_143, x_143); -x_232 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__74; +x_232 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__75; lean_inc(x_9); x_233 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_233, 0, x_9); lean_ctor_set(x_233, 1, x_232); -x_234 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__39; +x_234 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__40; lean_inc(x_143); lean_inc(x_197); lean_inc(x_182); lean_inc(x_9); x_235 = l_Lean_Syntax_node4(x_9, x_234, x_182, x_223, x_197, x_143); -x_236 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__75; +x_236 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__76; lean_inc_n(x_143, 2); lean_inc(x_9); x_237 = l_Lean_Syntax_node6(x_9, x_236, x_201, x_233, x_143, x_143, x_226, x_235); @@ -1448,14 +1457,14 @@ x_248 = lean_ctor_get(x_195, 0); lean_inc(x_248); lean_dec(x_195); x_249 = l_String_intercalate(x_183, x_248); -x_250 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__78; +x_250 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__79; x_251 = lean_string_append(x_250, x_249); lean_dec(x_249); x_252 = lean_box(2); x_253 = l_Lean_Syntax_mkNameLit(x_251, x_252); -x_254 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__79; +x_254 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__80; x_255 = lean_array_push(x_254, x_253); -x_256 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__77; +x_256 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__78; x_257 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_257, 0, x_252); lean_ctor_set(x_257, 1, x_256); @@ -1918,7 +1927,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_TypeName__ _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__30; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__31; x_3 = l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_TypeName___hyg_685____closed__1; x_4 = l_Lean_Elab_registerDerivingHandler(x_2, x_3, x_1); return x_4; @@ -2091,6 +2100,8 @@ l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__78); l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__79 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__79(); lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__79); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__80 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__80(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__1___closed__80); l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__2___closed__1 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__2___closed__1(); lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___lambda__2___closed__1); l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___closed__1 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__1___closed__1(); diff --git a/stage0/stdlib/Lean/Elab/Inductive.c b/stage0/stdlib/Lean/Elab/Inductive.c index 1f240a9c6eed..e955c805e99d 100644 --- a/stage0/stdlib/Lean/Elab/Inductive.c +++ b/stage0/stdlib/Lean/Elab/Inductive.c @@ -310,6 +310,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_ge static lean_object* l_Lean_Elab_Command_accLevelAtCtor___closed__4; uint8_t l_Array_contains___at_Lean_Meta_setMVarUserNamesAt___spec__1(lean_object*, lean_object*); lean_object* lean_st_ref_take(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__25; static lean_object* l_List_mapTR_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_fixedIndicesToParams___spec__5___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_fixedIndicesToParams___spec__6___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_fixedIndicesToParams___spec__7___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_elabHeaderAux(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_eq(lean_object*, lean_object*); @@ -32285,7 +32286,7 @@ static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Indu _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("def", 3); +x_1 = lean_mk_string_from_bytes("definition", 10); return x_1; } } @@ -32305,23 +32306,31 @@ static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Indu _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("declId", 6); +x_1 = lean_mk_string_from_bytes("def", 3); return x_1; } } static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__7() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("declId", 6); +return x_1; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__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_Lean_Elab_Command_initFn____x40_Lean_Elab_Inductive___hyg_5____closed__4; x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__1; x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Inductive___hyg_5____closed__7; -x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__6; +x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__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_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__8() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__9() { _start: { lean_object* x_1; @@ -32329,17 +32338,17 @@ x_1 = lean_mk_string_from_bytes("_root_", 6); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__9() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__8; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__9; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__10() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__11() { _start: { lean_object* x_1; @@ -32347,17 +32356,17 @@ x_1 = lean_mk_string_from_bytes("null", 4); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__11() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__10; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__11; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__12() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__13() { _start: { lean_object* x_1; @@ -32365,19 +32374,19 @@ x_1 = lean_mk_string_from_bytes("optDeclSig", 10); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__13() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___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_Elab_Command_initFn____x40_Lean_Elab_Inductive___hyg_5____closed__4; x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__1; x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Inductive___hyg_5____closed__7; -x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__12; +x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__13; 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___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__14() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__15() { _start: { lean_object* x_1; @@ -32385,7 +32394,7 @@ x_1 = lean_mk_string_from_bytes("Term", 4); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__15() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__16() { _start: { lean_object* x_1; @@ -32393,19 +32402,19 @@ x_1 = lean_mk_string_from_bytes("typeSpec", 8); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__16() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___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_Elab_Command_initFn____x40_Lean_Elab_Inductive___hyg_5____closed__4; x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__1; -x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__14; -x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__15; +x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__15; +x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__16; 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___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__17() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__18() { _start: { lean_object* x_1; @@ -32413,7 +32422,7 @@ x_1 = lean_mk_string_from_bytes(":", 1); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__18() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__19() { _start: { lean_object* x_1; @@ -32421,19 +32430,19 @@ x_1 = lean_mk_string_from_bytes("declValEqns", 11); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__19() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__20() { _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_initFn____x40_Lean_Elab_Inductive___hyg_5____closed__4; x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__1; x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Inductive___hyg_5____closed__7; -x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__18; +x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__19; 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___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__20() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__21() { _start: { lean_object* x_1; @@ -32441,19 +32450,19 @@ x_1 = lean_mk_string_from_bytes("matchAltsWhereDecls", 19); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__21() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__22() { _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_initFn____x40_Lean_Elab_Inductive___hyg_5____closed__4; x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__1; -x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__14; -x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__20; +x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__15; +x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__21; 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___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__22() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__23() { _start: { lean_object* x_1; @@ -32461,7 +32470,7 @@ x_1 = lean_mk_string_from_bytes("Termination", 11); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__23() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__24() { _start: { lean_object* x_1; @@ -32469,14 +32478,14 @@ x_1 = lean_mk_string_from_bytes("suffix", 6); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__24() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__25() { _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_initFn____x40_Lean_Elab_Inductive___hyg_5____closed__4; x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__1; -x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__22; -x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__23; +x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__23; +x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__24; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } @@ -32506,48 +32515,48 @@ x_19 = lean_ctor_get(x_17, 0); lean_dec(x_19); x_20 = 1; x_21 = l_Lean_SourceInfo_fromRef(x_1, x_20); -x_22 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__4; +x_22 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__6; 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_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__9; +x_24 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__10; x_25 = l_Lean_Name_append(x_24, x_2); x_26 = l_Lean_Name_append(x_25, x_3); x_27 = lean_mk_syntax_ident(x_26); -x_28 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__11; +x_28 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__12; x_29 = l_Lean_Elab_Command_instInhabitedCtorView___closed__1; lean_inc(x_14); x_30 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_30, 0, x_14); lean_ctor_set(x_30, 1, x_28); lean_ctor_set(x_30, 2, x_29); -x_31 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__7; +x_31 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__8; lean_inc(x_30); lean_inc(x_14); x_32 = l_Lean_Syntax_node2(x_14, x_31, x_27, x_30); -x_33 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__17; +x_33 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__18; lean_inc(x_14); x_34 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_34, 0, x_14); lean_ctor_set(x_34, 1, x_33); -x_35 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__16; +x_35 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__17; lean_inc(x_14); x_36 = l_Lean_Syntax_node2(x_14, x_35, x_34, x_4); lean_inc(x_14); x_37 = l_Lean_Syntax_node1(x_14, x_28, x_36); -x_38 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__13; +x_38 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__14; lean_inc(x_30); lean_inc(x_14); x_39 = l_Lean_Syntax_node2(x_14, x_38, x_30, x_37); -x_40 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__24; +x_40 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__25; lean_inc_n(x_30, 2); lean_inc(x_14); x_41 = l_Lean_Syntax_node2(x_14, x_40, x_30, x_30); -x_42 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__21; +x_42 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__22; lean_inc(x_30); lean_inc(x_14); x_43 = l_Lean_Syntax_node3(x_14, x_42, x_5, x_41, x_30); -x_44 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__19; +x_44 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__20; lean_inc(x_14); x_45 = l_Lean_Syntax_node1(x_14, x_44, x_43); x_46 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__5; @@ -32566,48 +32575,48 @@ lean_inc(x_50); lean_dec(x_17); x_51 = 1; x_52 = l_Lean_SourceInfo_fromRef(x_1, x_51); -x_53 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__4; +x_53 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__6; x_54 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_54, 0, x_52); lean_ctor_set(x_54, 1, x_53); -x_55 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__9; +x_55 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__10; x_56 = l_Lean_Name_append(x_55, x_2); x_57 = l_Lean_Name_append(x_56, x_3); x_58 = lean_mk_syntax_ident(x_57); -x_59 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__11; +x_59 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__12; x_60 = l_Lean_Elab_Command_instInhabitedCtorView___closed__1; lean_inc(x_14); x_61 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_61, 0, x_14); lean_ctor_set(x_61, 1, x_59); lean_ctor_set(x_61, 2, x_60); -x_62 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__7; +x_62 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__8; lean_inc(x_61); lean_inc(x_14); x_63 = l_Lean_Syntax_node2(x_14, x_62, x_58, x_61); -x_64 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__17; +x_64 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__18; lean_inc(x_14); x_65 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_65, 0, x_14); lean_ctor_set(x_65, 1, x_64); -x_66 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__16; +x_66 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__17; lean_inc(x_14); x_67 = l_Lean_Syntax_node2(x_14, x_66, x_65, x_4); lean_inc(x_14); x_68 = l_Lean_Syntax_node1(x_14, x_59, x_67); -x_69 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__13; +x_69 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__14; lean_inc(x_61); lean_inc(x_14); x_70 = l_Lean_Syntax_node2(x_14, x_69, x_61, x_68); -x_71 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__24; +x_71 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__25; lean_inc_n(x_61, 2); lean_inc(x_14); x_72 = l_Lean_Syntax_node2(x_14, x_71, x_61, x_61); -x_73 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__21; +x_73 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__22; lean_inc(x_61); lean_inc(x_14); x_74 = l_Lean_Syntax_node3(x_14, x_73, x_5, x_72, x_61); -x_75 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__19; +x_75 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__20; lean_inc(x_14); x_76 = l_Lean_Syntax_node1(x_14, x_75, x_74); x_77 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__5; @@ -32747,7 +32756,7 @@ x_38 = l_Lean_Elab_Command_getMainModule___rarg(x_9, x_37); x_39 = lean_ctor_get(x_38, 1); lean_inc(x_39); lean_dec(x_38); -x_40 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__11; +x_40 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__12; x_41 = l_Lean_Elab_Command_instInhabitedCtorView___closed__1; lean_inc(x_35); x_42 = lean_alloc_ctor(1, 3, 0); @@ -32835,7 +32844,7 @@ x_76 = l_Lean_Elab_Command_getMainModule___rarg(x_9, x_75); x_77 = lean_ctor_get(x_76, 1); lean_inc(x_77); lean_dec(x_76); -x_78 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__11; +x_78 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__12; x_79 = l_Lean_Elab_Command_instInhabitedCtorView___closed__1; lean_inc(x_73); x_80 = lean_alloc_ctor(1, 3, 0); @@ -32884,7 +32893,7 @@ lean_ctor_set(x_98, 1, x_97); x_99 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__2___closed__5; lean_inc(x_92); x_100 = l_Lean_Syntax_node1(x_92, x_99, x_98); -x_101 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__11; +x_101 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__12; lean_inc(x_92); x_102 = l_Lean_Syntax_node1(x_92, x_101, x_100); x_103 = l_Lean_Elab_Command_instInhabitedCtorView___closed__1; @@ -33080,7 +33089,7 @@ x_39 = l_Lean_Elab_Command_getMainModule___rarg(x_9, x_38); x_40 = lean_ctor_get(x_39, 1); lean_inc(x_40); lean_dec(x_39); -x_41 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__11; +x_41 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__12; x_42 = l_Lean_Elab_Command_instInhabitedCtorView___closed__1; lean_inc(x_36); x_43 = lean_alloc_ctor(1, 3, 0); @@ -33170,7 +33179,7 @@ x_78 = l_Lean_Elab_Command_getMainModule___rarg(x_9, x_77); x_79 = lean_ctor_get(x_78, 1); lean_inc(x_79); lean_dec(x_78); -x_80 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__11; +x_80 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__12; x_81 = l_Lean_Elab_Command_instInhabitedCtorView___closed__1; lean_inc(x_75); x_82 = lean_alloc_ctor(1, 3, 0); @@ -33284,7 +33293,7 @@ x_38 = l_Lean_Elab_Command_getMainModule___rarg(x_8, x_37); x_39 = lean_ctor_get(x_38, 1); lean_inc(x_39); lean_dec(x_38); -x_40 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__11; +x_40 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__12; x_41 = l_Lean_Elab_Command_instInhabitedCtorView___closed__1; lean_inc(x_35); x_42 = lean_alloc_ctor(1, 3, 0); @@ -33344,7 +33353,7 @@ static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Indu 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_initFn____x40_Lean_Elab_Inductive___hyg_5____closed__4; x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__1; -x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__14; +x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__15; x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__5___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -33422,7 +33431,7 @@ x_36 = l_Lean_Elab_Command_getMainModule___rarg(x_7, x_35); x_37 = lean_ctor_get(x_36, 1); lean_inc(x_37); lean_dec(x_36); -x_38 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__11; +x_38 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__12; x_39 = l_Lean_Elab_Command_instInhabitedCtorView___closed__1; lean_inc(x_33); x_40 = lean_alloc_ctor(1, 3, 0); @@ -33512,7 +33521,7 @@ x_75 = l_Lean_Elab_Command_getMainModule___rarg(x_7, x_74); x_76 = lean_ctor_get(x_75, 1); lean_inc(x_76); lean_dec(x_75); -x_77 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__11; +x_77 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__12; x_78 = l_Lean_Elab_Command_instInhabitedCtorView___closed__1; lean_inc(x_72); x_79 = lean_alloc_ctor(1, 3, 0); @@ -33727,7 +33736,7 @@ x_51 = l_Lean_Elab_Command_getMainModule___rarg(x_7, x_50); x_52 = lean_ctor_get(x_51, 1); lean_inc(x_52); lean_dec(x_51); -x_53 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__11; +x_53 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__12; x_54 = l_Lean_Elab_Command_instInhabitedCtorView___closed__1; lean_inc(x_48); x_55 = lean_alloc_ctor(1, 3, 0); @@ -33829,7 +33838,7 @@ x_93 = l_Lean_Elab_Command_getMainModule___rarg(x_7, x_92); x_94 = lean_ctor_get(x_93, 1); lean_inc(x_94); lean_dec(x_93); -x_95 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__11; +x_95 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__12; x_96 = l_Lean_Elab_Command_instInhabitedCtorView___closed__1; lean_inc(x_90); x_97 = lean_alloc_ctor(1, 3, 0); @@ -33927,7 +33936,7 @@ x_133 = l_Lean_Elab_Command_getMainModule___rarg(x_7, x_132); x_134 = lean_ctor_get(x_133, 1); lean_inc(x_134); lean_dec(x_133); -x_135 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__11; +x_135 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__12; x_136 = l_Lean_Elab_Command_instInhabitedCtorView___closed__1; lean_inc(x_130); x_137 = lean_alloc_ctor(1, 3, 0); @@ -34425,7 +34434,7 @@ lean_ctor_set(x_25, 0, x_19); lean_ctor_set(x_25, 1, x_24); x_26 = l_Lean_Elab_Command_instInhabitedCtorView___closed__1; x_27 = l_Array_append___rarg(x_26, x_13); -x_28 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__11; +x_28 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__12; lean_inc(x_19); x_29 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_29, 0, x_19); @@ -36240,6 +36249,8 @@ l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__23); l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__24 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__24(); lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__24); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__25 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__25(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__25); l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__2___closed__1 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__2___closed__1(); lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__2___closed__1); l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__2___closed__2 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__2___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/Eqns.c b/stage0/stdlib/Lean/Elab/PreDefinition/Eqns.c index 00ad068bd885..bc5dfe4f4559 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/Eqns.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/Eqns.c @@ -24771,7 +24771,7 @@ static lean_object* _init_l_Lean_Elab_Eqns_mkUnfoldEq___lambda__1___closed__1() _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("_unfold", 7); +x_1 = lean_mk_string_from_bytes("def", 3); return x_1; } } diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/Structural/Eqns.c b/stage0/stdlib/Lean/Elab/PreDefinition/Structural/Eqns.c index 4517ea83f29b..2b7711bebc88 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/Structural/Eqns.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/Structural/Eqns.c @@ -1742,7 +1742,7 @@ static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Structural_mkEqn _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("_eq", 3); +x_1 = lean_mk_string_from_bytes("eq", 2); return x_1; } } diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/WF/Eqns.c b/stage0/stdlib/Lean/Elab/PreDefinition/WF/Eqns.c index 14170b3947d8..886e4820442e 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/WF/Eqns.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/WF/Eqns.c @@ -4765,7 +4765,7 @@ static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_WF_mkEqns___spec _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("_eq", 3); +x_1 = lean_mk_string_from_bytes("eq", 2); return x_1; } } diff --git a/stage0/stdlib/Lean/Elab/Syntax.c b/stage0/stdlib/Lean/Elab/Syntax.c index f68494a6a399..a6081dc64575 100644 --- a/stage0/stdlib/Lean/Elab/Syntax.c +++ b/stage0/stdlib/Lean/Elab/Syntax.c @@ -79,6 +79,7 @@ static lean_object* l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__ static lean_object* l_Lean_Elab_Term_toParserDescr_processAlias___lambda__2___closed__11; extern lean_object* l_Lean_maxRecDepthErrorMessage; static lean_object* l___regBuiltin_Lean_Elab_Command_elabDeclareSyntaxCat_declRange___closed__2; +static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__74; static lean_object* l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__7; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_elabParserName_x3f___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_checkLeftRec___closed__1; @@ -11506,7 +11507,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("def", 3); +x_1 = lean_mk_string_from_bytes("definition", 10); return x_1; } } @@ -11526,23 +11527,31 @@ 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("declId", 6); +x_1 = lean_mk_string_from_bytes("def", 3); return x_1; } } static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__27() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("declId", 6); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__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_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__1; x_2 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__2; x_3 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__5; -x_4 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__26; +x_4 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__27; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__28() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__29() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; @@ -11556,7 +11565,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__29() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__30() { _start: { lean_object* x_1; @@ -11564,19 +11573,19 @@ x_1 = lean_mk_string_from_bytes("optDeclSig", 10); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__30() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__31() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__1; x_2 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__2; x_3 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__5; -x_4 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__29; +x_4 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__30; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__31() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__32() { _start: { lean_object* x_1; @@ -11584,19 +11593,19 @@ x_1 = lean_mk_string_from_bytes("typeSpec", 8); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__32() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___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_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__1; x_2 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__2; x_3 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__3; -x_4 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__31; +x_4 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__32; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__33() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__34() { _start: { lean_object* x_1; @@ -11604,7 +11613,7 @@ x_1 = lean_mk_string_from_bytes(":", 1); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__34() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__35() { _start: { lean_object* x_1; @@ -11612,16 +11621,16 @@ x_1 = lean_mk_string_from_bytes("Lean.ParserDescr", 16); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__35() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__36() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__34; +x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__35; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__36() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__37() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -11631,53 +11640,53 @@ x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__37() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__38() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__36; +x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__37; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__38() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__39() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__36; +x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__37; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__39() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__40() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__38; +x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__39; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__40() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__41() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__37; -x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__39; +x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__38; +x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__40; 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___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__41() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__42() { _start: { lean_object* x_1; @@ -11685,19 +11694,19 @@ x_1 = lean_mk_string_from_bytes("declValSimple", 13); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__42() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__43() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__1; x_2 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__2; x_3 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__5; -x_4 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__41; +x_4 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__42; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__43() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__44() { _start: { lean_object* x_1; @@ -11705,7 +11714,7 @@ x_1 = lean_mk_string_from_bytes(":=", 2); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__44() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__45() { _start: { lean_object* x_1; @@ -11713,16 +11722,16 @@ x_1 = lean_mk_string_from_bytes("Lean.ParserDescr.node", 21); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__45() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__46() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__44; +x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__45; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__46() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__47() { _start: { lean_object* x_1; @@ -11730,64 +11739,64 @@ x_1 = lean_mk_string_from_bytes("node", 4); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__47() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__48() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__1; x_2 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; -x_3 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__46; +x_3 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__47; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__48() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__49() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__47; +x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__48; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__49() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__50() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__47; +x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__48; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__50() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__51() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__49; +x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__50; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__51() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__52() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__48; -x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__50; +x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__49; +x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__51; 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___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__52() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__53() { _start: { lean_object* x_1; @@ -11795,7 +11804,7 @@ x_1 = lean_mk_string_from_bytes("`Lean.Parser.Term.quot", 22); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__53() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__54() { _start: { lean_object* x_1; @@ -11803,17 +11812,17 @@ x_1 = lean_mk_string_from_bytes("num", 3); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__54() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__55() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__53; +x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__54; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__55() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__56() { _start: { lean_object* x_1; lean_object* x_2; @@ -11822,17 +11831,17 @@ x_2 = l_Nat_repr(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__56() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__57() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__55; +x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__56; x_2 = lean_box(2); x_3 = l_Lean_Syntax_mkNumLit(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__57() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__58() { _start: { lean_object* x_1; @@ -11840,16 +11849,16 @@ x_1 = lean_mk_string_from_bytes("Lean.ParserDescr.binary", 23); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__58() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__59() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__57; +x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__58; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__59() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__60() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -11861,11 +11870,11 @@ lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__60() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__61() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__59; +x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__60; x_2 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__14; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11873,7 +11882,7 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__61() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__62() { _start: { lean_object* x_1; @@ -11881,16 +11890,16 @@ x_1 = lean_mk_string_from_bytes("Lean.ParserDescr.symbol", 23); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__62() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__63() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__61; +x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__62; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__63() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__64() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -11902,11 +11911,11 @@ lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__64() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__65() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__63; +x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__64; x_2 = l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__8; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11914,7 +11923,7 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__65() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__66() { _start: { lean_object* x_1; @@ -11922,16 +11931,16 @@ x_1 = lean_mk_string_from_bytes("Lean.ParserDescr.cat", 20); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__66() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__67() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__65; +x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__66; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__67() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__68() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -11943,11 +11952,11 @@ lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__68() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__69() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__67; +x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__68; x_2 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__7; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11955,7 +11964,7 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__69() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__70() { _start: { lean_object* x_1; @@ -11963,7 +11972,7 @@ x_1 = lean_mk_string_from_bytes("0", 1); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__70() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__71() { _start: { lean_object* x_1; @@ -11971,7 +11980,7 @@ x_1 = lean_mk_string_from_bytes("\")\"", 3); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__71() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__72() { _start: { lean_object* x_1; @@ -11979,7 +11988,7 @@ x_1 = lean_mk_string_from_bytes("Termination", 11); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__72() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__73() { _start: { lean_object* x_1; @@ -11987,14 +11996,14 @@ x_1 = lean_mk_string_from_bytes("suffix", 6); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__73() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__74() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__1; x_2 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__2; -x_3 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__71; -x_4 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__72; +x_3 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__72; +x_4 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__73; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } @@ -12086,7 +12095,7 @@ x_45 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotPar lean_inc_n(x_25, 5); lean_inc(x_16); x_46 = l_Lean_Syntax_node6(x_16, x_45, x_25, x_44, x_25, x_25, x_25, x_25); -x_47 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__24; +x_47 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__26; lean_inc(x_16); x_48 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_48, 0, x_16); @@ -12095,58 +12104,58 @@ lean_inc(x_11); x_49 = lean_mk_syntax_ident(x_11); x_50 = l_Lean_Elab_Term_toParserDescr_process___closed__16; x_51 = lean_array_push(x_50, x_49); -x_52 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__28; +x_52 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__29; x_53 = lean_array_push(x_51, x_52); x_54 = lean_box(2); -x_55 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__27; +x_55 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__28; x_56 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_56, 0, x_54); lean_ctor_set(x_56, 1, x_55); lean_ctor_set(x_56, 2, x_53); -x_57 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__33; +x_57 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__34; lean_inc(x_16); x_58 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_58, 0, x_16); lean_ctor_set(x_58, 1, x_57); -x_59 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__36; +x_59 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__37; lean_inc(x_18); lean_inc(x_21); x_60 = l_Lean_addMacroScope(x_21, x_59, x_18); -x_61 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__35; -x_62 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__40; +x_61 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__36; +x_62 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__41; lean_inc(x_16); x_63 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_63, 0, x_16); lean_ctor_set(x_63, 1, x_61); lean_ctor_set(x_63, 2, x_60); lean_ctor_set(x_63, 3, x_62); -x_64 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__32; +x_64 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__33; lean_inc(x_16); x_65 = l_Lean_Syntax_node2(x_16, x_64, x_58, x_63); lean_inc(x_16); x_66 = l_Lean_Syntax_node1(x_16, x_23, x_65); -x_67 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__30; +x_67 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__31; lean_inc(x_25); lean_inc(x_16); x_68 = l_Lean_Syntax_node2(x_16, x_67, x_25, x_66); -x_69 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__43; +x_69 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__44; lean_inc(x_16); x_70 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_70, 0, x_16); lean_ctor_set(x_70, 1, x_69); -x_71 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__47; +x_71 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__48; lean_inc(x_18); lean_inc(x_21); x_72 = l_Lean_addMacroScope(x_21, x_71, x_18); -x_73 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__45; -x_74 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__51; +x_73 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__46; +x_74 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__52; lean_inc(x_16); x_75 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_75, 0, x_16); lean_ctor_set(x_75, 1, x_73); lean_ctor_set(x_75, 2, x_72); lean_ctor_set(x_75, 3, x_74); -x_76 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__52; +x_76 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__53; lean_inc(x_16); x_77 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_77, 0, x_16); @@ -12168,8 +12177,8 @@ x_85 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_ lean_inc(x_18); lean_inc(x_21); x_86 = l_Lean_addMacroScope(x_21, x_85, x_18); -x_87 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__58; -x_88 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__60; +x_87 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__59; +x_88 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__61; lean_inc(x_16); x_89 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_89, 0, x_16); @@ -12189,8 +12198,8 @@ x_94 = l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__5; lean_inc(x_18); lean_inc(x_21); x_95 = l_Lean_addMacroScope(x_21, x_94, x_18); -x_96 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__62; -x_97 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__64; +x_96 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__63; +x_97 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__65; lean_inc(x_16); x_98 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_98, 0, x_16); @@ -12217,8 +12226,8 @@ lean_inc(x_16); x_106 = l_Lean_Syntax_node3(x_16, x_105, x_83, x_102, x_104); x_107 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__4; x_108 = l_Lean_addMacroScope(x_21, x_107, x_18); -x_109 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__66; -x_110 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__68; +x_109 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__67; +x_110 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__69; lean_inc(x_16); x_111 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_111, 0, x_16); @@ -12227,15 +12236,15 @@ lean_ctor_set(x_111, 2, x_108); lean_ctor_set(x_111, 3, x_110); lean_inc(x_1); x_112 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_32, x_1); -x_113 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__69; +x_113 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__70; lean_inc(x_16); x_114 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_114, 0, x_16); lean_ctor_set(x_114, 1, x_113); -x_115 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__54; +x_115 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__55; lean_inc(x_16); x_116 = l_Lean_Syntax_node1(x_16, x_115, x_114); -x_117 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__70; +x_117 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__71; lean_inc(x_16); x_118 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_118, 0, x_16); @@ -12251,7 +12260,7 @@ lean_inc(x_104); lean_inc(x_83); lean_inc(x_16); x_123 = l_Lean_Syntax_node3(x_16, x_105, x_83, x_122, x_104); -x_124 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__73; +x_124 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__74; lean_inc_n(x_25, 2); lean_inc(x_16); x_125 = l_Lean_Syntax_node2(x_16, x_124, x_25, x_25); @@ -12316,7 +12325,7 @@ lean_inc(x_104); lean_inc(x_83); lean_inc(x_16); x_136 = l_Lean_Syntax_node3(x_16, x_105, x_83, x_135, x_104); -x_137 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__56; +x_137 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__57; lean_inc(x_16); x_138 = l_Lean_Syntax_node3(x_16, x_23, x_126, x_137, x_136); lean_inc(x_75); @@ -12328,7 +12337,7 @@ lean_inc(x_16); x_141 = l_Lean_Syntax_node3(x_16, x_23, x_81, x_137, x_140); lean_inc(x_16); x_142 = l_Lean_Syntax_node2(x_16, x_101, x_75, x_141); -x_143 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__42; +x_143 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__43; lean_inc(x_25); lean_inc(x_16); x_144 = l_Lean_Syntax_node4(x_16, x_143, x_70, x_142, x_125, x_25); @@ -12385,7 +12394,7 @@ lean_inc(x_104); lean_inc(x_83); lean_inc(x_16); x_167 = l_Lean_Syntax_node3(x_16, x_105, x_83, x_166, x_104); -x_168 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__56; +x_168 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__57; lean_inc(x_16); x_169 = l_Lean_Syntax_node3(x_16, x_23, x_126, x_168, x_167); lean_inc(x_75); @@ -12397,7 +12406,7 @@ lean_inc(x_16); x_172 = l_Lean_Syntax_node3(x_16, x_23, x_81, x_168, x_171); lean_inc(x_16); x_173 = l_Lean_Syntax_node2(x_16, x_101, x_75, x_172); -x_174 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__42; +x_174 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__43; lean_inc(x_25); lean_inc(x_16); x_175 = l_Lean_Syntax_node4(x_16, x_174, x_70, x_173, x_125, x_25); @@ -12848,7 +12857,7 @@ x_42 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_42, 0, x_33); lean_ctor_set(x_42, 1, x_40); lean_ctor_set(x_42, 2, x_41); -x_43 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__24; +x_43 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__26; lean_inc(x_33); x_44 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_44, 0, x_33); @@ -12859,15 +12868,15 @@ x_47 = 1; x_48 = l_Lean_mkIdentFrom(x_9, x_46, x_47); x_49 = l_Lean_Elab_Term_toParserDescr_process___closed__16; x_50 = lean_array_push(x_49, x_48); -x_51 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__28; +x_51 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__29; x_52 = lean_array_push(x_50, x_51); x_53 = lean_box(2); -x_54 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__27; +x_54 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__28; x_55 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_55, 0, x_53); lean_ctor_set(x_55, 1, x_54); lean_ctor_set(x_55, 2, x_52); -x_56 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__33; +x_56 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__34; lean_inc(x_33); x_57 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_57, 0, x_33); @@ -12881,16 +12890,16 @@ lean_ctor_set(x_61, 0, x_33); lean_ctor_set(x_61, 1, x_59); lean_ctor_set(x_61, 2, x_58); lean_ctor_set(x_61, 3, x_60); -x_62 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__32; +x_62 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__33; lean_inc(x_33); x_63 = l_Lean_Syntax_node2(x_33, x_62, x_57, x_61); lean_inc(x_33); x_64 = l_Lean_Syntax_node1(x_33, x_40, x_63); -x_65 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__30; +x_65 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__31; lean_inc(x_42); lean_inc(x_33); x_66 = l_Lean_Syntax_node2(x_33, x_65, x_42, x_64); -x_67 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__43; +x_67 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__44; lean_inc(x_33); x_68 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_68, 0, x_33); @@ -12921,11 +12930,11 @@ x_78 = l_Lean_Syntax_node6(x_33, x_77, x_70, x_42, x_42, x_76, x_42, x_72); x_79 = l_Lean_Elab_Term_toParserDescr_process___closed__2; lean_inc(x_33); x_80 = l_Lean_Syntax_node2(x_33, x_79, x_74, x_78); -x_81 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__73; +x_81 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__74; lean_inc_n(x_42, 2); lean_inc(x_33); x_82 = l_Lean_Syntax_node2(x_33, x_81, x_42, x_42); -x_83 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__42; +x_83 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__43; lean_inc(x_42); lean_inc(x_33); x_84 = l_Lean_Syntax_node4(x_33, x_83, x_68, x_80, x_82, x_42); @@ -15086,7 +15095,7 @@ static lean_object* _init_l_Lean_Elab_Command_elabSyntax___lambda__4___closed__2 { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__36; +x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__37; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); @@ -15098,7 +15107,7 @@ static lean_object* _init_l_Lean_Elab_Command_elabSyntax___lambda__4___closed__3 { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Command_elabSyntax___lambda__4___closed__2; -x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__39; +x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__40; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -15127,7 +15136,7 @@ static lean_object* _init_l_Lean_Elab_Command_elabSyntax___lambda__4___closed__6 { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; -x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__46; +x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__47; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -15472,26 +15481,26 @@ x_76 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_76, 0, x_59); lean_ctor_set(x_76, 1, x_4); lean_ctor_set(x_76, 2, x_68); -x_77 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__24; +x_77 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__26; lean_inc(x_59); x_78 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_78, 0, x_59); lean_ctor_set(x_78, 1, x_77); -x_79 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__27; +x_79 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__28; lean_inc(x_76); lean_inc(x_59); x_80 = l_Lean_Syntax_node2(x_59, x_79, x_36, x_76); -x_81 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__33; +x_81 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__34; lean_inc(x_59); x_82 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_82, 0, x_59); lean_ctor_set(x_82, 1, x_81); -x_83 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__36; +x_83 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__37; lean_inc(x_61); lean_inc(x_64); x_84 = l_Lean_addMacroScope(x_64, x_83, x_61); x_85 = lean_box(0); -x_86 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__35; +x_86 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__36; x_87 = l_Lean_Elab_Command_elabSyntax___lambda__4___closed__3; lean_inc(x_59); x_88 = lean_alloc_ctor(3, 4, 0); @@ -15499,17 +15508,17 @@ lean_ctor_set(x_88, 0, x_59); lean_ctor_set(x_88, 1, x_86); lean_ctor_set(x_88, 2, x_84); lean_ctor_set(x_88, 3, x_87); -x_89 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__32; +x_89 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__33; lean_inc(x_59); x_90 = l_Lean_Syntax_node2(x_59, x_89, x_82, x_88); lean_inc(x_4); lean_inc(x_59); x_91 = l_Lean_Syntax_node1(x_59, x_4, x_90); -x_92 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__30; +x_92 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__31; lean_inc(x_76); lean_inc(x_59); x_93 = l_Lean_Syntax_node2(x_59, x_92, x_76, x_91); -x_94 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__43; +x_94 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__44; lean_inc(x_59); x_95 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_95, 0, x_59); @@ -15517,7 +15526,7 @@ lean_ctor_set(x_95, 1, x_94); x_96 = l_Lean_Elab_Command_elabSyntax___lambda__4___closed__6; x_97 = l_Lean_addMacroScope(x_64, x_96, x_61); x_98 = l_Lean_Elab_Command_elabSyntax___lambda__4___closed__5; -x_99 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__51; +x_99 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__52; lean_inc(x_59); x_100 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_100, 0, x_59); @@ -15528,7 +15537,7 @@ lean_inc(x_25); x_101 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_85, x_25); x_102 = l_Nat_repr(x_7); x_103 = l_Lean_Syntax_mkNumLit(x_102, x_46); -x_104 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__73; +x_104 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__74; lean_inc_n(x_76, 2); lean_inc(x_59); x_105 = l_Lean_Syntax_node2(x_59, x_104, x_76, x_76); @@ -15570,7 +15579,7 @@ x_112 = l_Lean_Syntax_node3(x_59, x_4, x_111, x_103, x_34); x_113 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__5; lean_inc(x_59); x_114 = l_Lean_Syntax_node2(x_59, x_113, x_100, x_112); -x_115 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__42; +x_115 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__43; lean_inc(x_76); lean_inc(x_59); x_116 = l_Lean_Syntax_node4(x_59, x_115, x_95, x_114, x_105, x_76); @@ -15607,7 +15616,7 @@ x_132 = l_Lean_Syntax_node3(x_59, x_4, x_131, x_103, x_34); x_133 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__5; lean_inc(x_59); x_134 = l_Lean_Syntax_node2(x_59, x_133, x_100, x_132); -x_135 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__42; +x_135 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__43; lean_inc(x_76); lean_inc(x_59); x_136 = l_Lean_Syntax_node4(x_59, x_135, x_95, x_134, x_105, x_76); @@ -15676,16 +15685,16 @@ x_166 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_166, 0, x_149); lean_ctor_set(x_166, 1, x_4); lean_ctor_set(x_166, 2, x_158); -x_167 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__24; +x_167 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__26; lean_inc(x_149); x_168 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_168, 0, x_149); lean_ctor_set(x_168, 1, x_167); -x_169 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__27; +x_169 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__28; lean_inc(x_166); lean_inc(x_149); x_170 = l_Lean_Syntax_node2(x_149, x_169, x_36, x_166); -x_171 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__33; +x_171 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__34; lean_inc(x_149); x_172 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_172, 0, x_149); @@ -15703,17 +15712,17 @@ lean_ctor_set(x_178, 0, x_149); lean_ctor_set(x_178, 1, x_176); lean_ctor_set(x_178, 2, x_174); lean_ctor_set(x_178, 3, x_177); -x_179 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__32; +x_179 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__33; lean_inc(x_149); x_180 = l_Lean_Syntax_node2(x_149, x_179, x_172, x_178); lean_inc(x_4); lean_inc(x_149); x_181 = l_Lean_Syntax_node1(x_149, x_4, x_180); -x_182 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__30; +x_182 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__31; lean_inc(x_166); lean_inc(x_149); x_183 = l_Lean_Syntax_node2(x_149, x_182, x_166, x_181); -x_184 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__43; +x_184 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__44; lean_inc(x_149); x_185 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_185, 0, x_149); @@ -15734,7 +15743,7 @@ x_192 = l_Nat_repr(x_7); x_193 = l_Lean_Syntax_mkNumLit(x_192, x_46); x_194 = l_Nat_repr(x_145); x_195 = l_Lean_Syntax_mkNumLit(x_194, x_46); -x_196 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__73; +x_196 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__74; lean_inc_n(x_166, 2); lean_inc(x_149); x_197 = l_Lean_Syntax_node2(x_149, x_196, x_166, x_166); @@ -15776,7 +15785,7 @@ x_204 = l_Lean_Syntax_node4(x_149, x_4, x_203, x_193, x_195, x_34); x_205 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__5; lean_inc(x_149); x_206 = l_Lean_Syntax_node2(x_149, x_205, x_190, x_204); -x_207 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__42; +x_207 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__43; lean_inc(x_166); lean_inc(x_149); x_208 = l_Lean_Syntax_node4(x_149, x_207, x_185, x_206, x_197, x_166); @@ -15813,7 +15822,7 @@ x_224 = l_Lean_Syntax_node4(x_149, x_4, x_223, x_193, x_195, x_34); x_225 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__5; lean_inc(x_149); x_226 = l_Lean_Syntax_node2(x_149, x_225, x_190, x_224); -x_227 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__42; +x_227 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__43; lean_inc(x_166); lean_inc(x_149); x_228 = l_Lean_Syntax_node4(x_149, x_227, x_185, x_226, x_197, x_166); @@ -17224,26 +17233,26 @@ x_43 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_43, 0, x_34); lean_ctor_set(x_43, 1, x_41); lean_ctor_set(x_43, 2, x_42); -x_44 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__24; +x_44 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__26; lean_inc(x_34); x_45 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_45, 0, x_34); lean_ctor_set(x_45, 1, x_44); -x_46 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__27; +x_46 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__28; lean_inc(x_43); lean_inc(x_34); x_47 = l_Lean_Syntax_node2(x_34, x_46, x_8, x_43); -x_48 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__33; +x_48 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__34; lean_inc(x_34); x_49 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_49, 0, x_34); lean_ctor_set(x_49, 1, x_48); -x_50 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__36; +x_50 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__37; lean_inc(x_36); lean_inc(x_39); x_51 = l_Lean_addMacroScope(x_39, x_50, x_36); x_52 = lean_box(0); -x_53 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__35; +x_53 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__36; x_54 = l_Lean_Elab_Command_elabSyntax___lambda__4___closed__3; lean_inc(x_34); x_55 = lean_alloc_ctor(3, 4, 0); @@ -17251,16 +17260,16 @@ lean_ctor_set(x_55, 0, x_34); lean_ctor_set(x_55, 1, x_53); lean_ctor_set(x_55, 2, x_51); lean_ctor_set(x_55, 3, x_54); -x_56 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__32; +x_56 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__33; lean_inc(x_34); x_57 = l_Lean_Syntax_node2(x_34, x_56, x_49, x_55); lean_inc(x_34); x_58 = l_Lean_Syntax_node1(x_34, x_41, x_57); -x_59 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__30; +x_59 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__31; lean_inc(x_43); lean_inc(x_34); x_60 = l_Lean_Syntax_node2(x_34, x_59, x_43, x_58); -x_61 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__43; +x_61 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__44; lean_inc(x_34); x_62 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_62, 0, x_34); @@ -17282,7 +17291,7 @@ x_71 = l_Lean_Syntax_mkStrLit(x_69, x_70); lean_dec(x_69); lean_inc(x_29); x_72 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_52, x_29); -x_73 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__73; +x_73 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__74; lean_inc_n(x_43, 2); lean_inc(x_34); x_74 = l_Lean_Syntax_node2(x_34, x_73, x_43, x_43); @@ -17323,7 +17332,7 @@ x_81 = l_Lean_Syntax_node3(x_34, x_41, x_71, x_80, x_23); x_82 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__5; lean_inc(x_34); x_83 = l_Lean_Syntax_node2(x_34, x_82, x_67, x_81); -x_84 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__42; +x_84 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__43; lean_inc(x_43); lean_inc(x_34); x_85 = l_Lean_Syntax_node4(x_34, x_84, x_62, x_83, x_74, x_43); @@ -17363,7 +17372,7 @@ x_102 = l_Lean_Syntax_node3(x_34, x_41, x_71, x_101, x_23); x_103 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__5; lean_inc(x_34); x_104 = l_Lean_Syntax_node2(x_34, x_103, x_67, x_102); -x_105 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__42; +x_105 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__43; lean_inc(x_43); lean_inc(x_34); x_106 = l_Lean_Syntax_node4(x_34, x_105, x_62, x_104, x_74, x_43); @@ -19474,6 +19483,8 @@ l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___c lean_mark_persistent(l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__72); l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__73 = _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__73(); lean_mark_persistent(l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__73); +l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__74 = _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__74(); +lean_mark_persistent(l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__74); l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__1 = _init_l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__1(); lean_mark_persistent(l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__1); l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__2 = _init_l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/Tactic/Omega/Core.c b/stage0/stdlib/Lean/Elab/Tactic/Omega/Core.c index efc231d30d99..68655faf9c74 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Omega/Core.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Omega/Core.c @@ -14,6 +14,7 @@ extern "C" { #endif LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Tactic_Omega_Problem_fourierMotzkin___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_elimination___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_Expr_const___override(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Elab_Tactic_Omega_Problem_possible___default; static lean_object* l_Lean_Elab_Tactic_Omega_Justification_toString___closed__15; @@ -28,7 +29,6 @@ static lean_object* l_Lean_Elab_Tactic_Omega_Problem_solveEasyEquality___closed_ static lean_object* l___auto____x40_Lean_Elab_Tactic_Omega_Core___hyg_1659____closed__15; static lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Core___hyg_5____closed__13; static lean_object* l___auto____x40_Lean_Elab_Tactic_Omega_Core___hyg_1659____closed__5; -lean_object* l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___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_Omega_instToExprLinearCombo___lambda__1___closed__11; lean_object* l_Lean_mkNatLit(lean_object*); LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Elab_Tactic_Omega_Problem_fourierMotzkinData___spec__1___boxed(lean_object*, lean_object*); @@ -47,11 +47,12 @@ uint8_t l___private_Init_Omega_Constraint_0__Lean_Omega_decEqConstraint____x40_I static lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1___closed__4; static lean_object* l_Lean_Elab_Tactic_Omega_instToExprConstraint___closed__4; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Tactic_Omega_Problem_fourierMotzkinData___spec__2___lambda__1___boxed(lean_object*, lean_object*); -lean_object* l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(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_Tactic_Omega_lookup___spec__14(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* l_Lean_Omega_Constraint_combo(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_Justification_comboProof___closed__2; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___boxed(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_Omega_Problem_dealWithHardEquality___closed__5; -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___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_throwError___at_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_Justification_tidyProof___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Justification_proof___boxed(lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_Problem_addInequality___closed__1; @@ -125,7 +126,7 @@ static lean_object* l_Lean_Elab_Tactic_Omega_Justification_bmodProof___closed__1 static lean_object* l_Lean_Elab_Tactic_Omega_Justification_bmodProof___closed__21; LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Tactic_Omega_Problem_fourierMotzkin___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_Elab_Tactic_Omega_Problem_FourierMotzkinData_lowerBounds___default; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_proveFalse___boxed(lean_object*, lean_object*, lean_object*, 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_Omega_Problem_proveFalse___boxed(lean_object*, lean_object*, 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_Omega_Justification_proof___rarg___closed__1; static lean_object* l_Lean_Elab_Tactic_Omega_Justification_proof___rarg___closed__2; LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Tactic_Omega_Problem_fourierMotzkin___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -134,7 +135,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_instToStringFourierMot static lean_object* l_Lean_Elab_Tactic_Omega_Problem_instToStringProblem___closed__2; static lean_object* l_List_mapTR_loop___at_Lean_Elab_Tactic_Omega_Problem_fourierMotzkinSelect___spec__3___closed__11; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Tactic_Omega_Problem_fourierMotzkinSelect___spec__1___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_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_addInequality___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_addInequality___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_fourierMotzkinSelect(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_instToExprConstraint___lambda__1___closed__7; LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Elab_Tactic_Omega_Problem_fourierMotzkinData___spec__1(lean_object*, lean_object*); @@ -152,7 +153,7 @@ LEAN_EXPORT uint8_t l_List_beq___at_Lean_Elab_Tactic_Omega_Justification_toStrin LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_solveEasyEquality___lambda__2___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Omega_Core_0__Lean_Elab_Tactic_Omega_Justification_bullet___boxed(lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_Problem_proveFalse___closed__2; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___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_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Elab_Tactic_Omega_Core___hyg_1659____closed__12; static lean_object* l_List_mapTR_loop___at_Lean_Elab_Tactic_Omega_Problem_fourierMotzkinSelect___spec__3___closed__8; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Tactic_Omega_Problem_fourierMotzkinData___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -173,10 +174,11 @@ LEAN_EXPORT lean_object* l_Lean_HashSetImp_moveEntries___at_Lean_Elab_Tactic_Ome LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Elab_Tactic_Omega_Problem_insertConstraint___spec__2___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Fact_instToStringFact(lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_Justification_tidyProof___closed__7; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_addInequality___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_Tactic_Omega_Problem_addInequality___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___lambda__1___closed__3; static lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Core___hyg_5____closed__12; lean_object* l_Lean_mkApp6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_runOmega___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_Name_mkStr3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_fourierMotzkinSelect___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_Justification_bmodProof___closed__5; @@ -185,7 +187,7 @@ static lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omeg size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_solveEasyEquality___lambda__1___boxed(lean_object*); LEAN_EXPORT uint8_t l_Lean_Elab_Tactic_Omega_Problem_FourierMotzkinData_upperExact___default; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_elimination(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_Omega_Problem_elimination(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_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Core___hyg_5_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1(lean_object*, lean_object*, lean_object*); lean_object* l_Int_sign(lean_object*); @@ -219,30 +221,32 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Omega_P LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_fourierMotzkinSelect___lambda__4___boxed(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_Elab_Tactic_Omega_Problem_runOmega___closed__3; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Justification_proof___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_Omega_Justification_proof___rarg(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*); static lean_object* l_Lean_Elab_Tactic_Omega_instToExprConstraint___closed__2; lean_object* l_List_range(lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_Problem_instToStringFourierMotzkinData___closed__2; static lean_object* l_Lean_Elab_Tactic_Omega_Justification_toString___closed__8; static lean_object* l_Lean_Elab_Tactic_Omega_Problem_addEquality__proof___closed__1; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_addInequality___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_addInequality___lambda__1___boxed(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_Tactic_Omega_Problem_solveEasyEquality___lambda__1(lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_Problem_instToStringFourierMotzkinData___closed__4; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Omega_Problem_selectEquality___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Tactic_Omega_Problem_fourierMotzkinSelect___spec__1___closed__1; static lean_object* l_List_mapTR_loop___at_Lean_Elab_Tactic_Omega_Problem_fourierMotzkinSelect___spec__3___closed__6; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Justification_proof___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_Lean_Elab_Tactic_Omega_Justification_proof___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*); static lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Elab_Tactic_Omega_Problem_constraints___default___spec__1(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_addInequality___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_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_addInequality___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_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__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_Lean_Elab_Tactic_Omega_Problem_solveEqualities___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_Omega_Justification_proof___rarg___closed__6; static lean_object* l_List_foldr___at_Lean_Elab_Tactic_Omega_Problem_replayEliminations___spec__1___closed__1; static lean_object* l_Lean_Elab_Tactic_Omega_Problem_proveFalse___closed__11; static lean_object* l_List_foldl___at_Lean_Elab_Tactic_Omega_Justification_toString___spec__2___closed__2; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Tactic_Omega_Problem_fourierMotzkinSelect___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_List_foldr___at_Lean_Elab_Tactic_Omega_Problem_replayEliminations___spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_fourierMotzkin(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_addEquality__proof___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_Omega_Problem_addEquality__proof___boxed(lean_object*, 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_Omega_Problem_proveFalse___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_fourierMotzkinData(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_numVars___default; @@ -273,7 +277,7 @@ static lean_object* l_Lean_Elab_Tactic_Omega_Justification_bmodProof___closed__2 LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Fact_combo(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Core___hyg_5____closed__16; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_replayEliminations(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_solveEquality(lean_object*, lean_object*, 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_Omega_Problem_solveEquality(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_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Tactic_Omega_Problem_fourierMotzkinData___spec__2___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Core___hyg_5____closed__2; extern lean_object* l_Lean_instInhabitedExpr; @@ -284,14 +288,15 @@ static lean_object* l_Lean_Elab_Tactic_Omega_Problem_instToStringFourierMotzkinD static lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Core___hyg_5____closed__4; lean_object* l_List_lengthTRAux___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___lambda__1___closed__4; -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___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_throwError___at_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___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* l_Lean_Omega_tidy_x3f(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality(lean_object*, 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_Omega_Problem_dealWithHardEquality(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_instToExprConstraint___lambda__1___closed__4; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Omega_Problem_solveEasyEquality___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Tactic_Omega_Problem_fourierMotzkinData___spec__2___lambda__1(lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_Justification_proof___rarg___closed__3; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_elimination___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_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Omega_Problem_selectEquality___spec__2(lean_object*, size_t, size_t, lean_object*); lean_object* l_Lean_Elab_Tactic_Omega_mkEqReflWithExpectedType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_instToExprConstraint___closed__1; @@ -307,7 +312,7 @@ static lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Ela 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_Elab_Tactic_Omega_Core_0__Lean_Elab_Tactic_Omega_Justification_bullet(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_insertConstraint___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Elab_Tactic_Omega_atomsCoeffs___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Elab_Tactic_Omega_atomsCoeffs___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_Problem_proveFalse___closed__8; static lean_object* l_Lean_Elab_Tactic_Omega_Justification_toString___closed__5; static lean_object* l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___closed__4; @@ -323,10 +328,10 @@ static lean_object* l_List_mapTR_loop___at_Lean_Elab_Tactic_Omega_Problem_fourie static lean_object* l_Lean_Elab_Tactic_Omega_Justification_toString___closed__16; lean_object* l_Int_bmod(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_Justification_bmodProof___closed__10; -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___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_throwError___at_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___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_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Core___hyg_5____closed__14; LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Elab_Tactic_Omega_Problem_insertConstraint___spec__6(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_addEquality__proof(lean_object*, lean_object*, 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_Omega_Problem_addEquality__proof(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*); size_t lean_hashmap_mk_idx(lean_object*, uint64_t); lean_object* lean_mk_thunk(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Justification_comboProof___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -334,11 +339,12 @@ static lean_object* l_List_mapTR_loop___at_Lean_Elab_Tactic_Omega_Problem_fourie LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_fourierMotzkinSelect___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*); static lean_object* l___private_Lean_Elab_Tactic_Omega_Core_0__Lean_Elab_Tactic_Omega_Justification_bullet___closed__3; static lean_object* l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___lambda__1___closed__12; +lean_object* l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___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*); static lean_object* l_Lean_Elab_Tactic_Omega_Problem_solveEasyEquality___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_insertConstraint(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Justification_tidyProof___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Elab_Tactic_Omega_lookup(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_Omega_Problem_runOmega(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Elab_Tactic_Omega_lookup(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_Lean_Elab_Tactic_Omega_Problem_runOmega(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static uint8_t l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Omega_Problem_fourierMotzkinSelect___spec__4___closed__2; static lean_object* l_Lean_Elab_Tactic_Omega_Problem_addEquality___closed__2; static lean_object* l_Lean_Elab_Tactic_Omega_Justification_toString___closed__4; @@ -347,13 +353,13 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_instToExprConstraint___lambda_ static lean_object* l_Lean_Elab_Tactic_Omega_Justification_bmodProof___closed__9; static lean_object* l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___lambda__1___closed__10; static lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Core___hyg_5____closed__23; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_elimination___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_Omega_Problem_elimination___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*); static lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Core___hyg_5____closed__20; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_instToExprConstraint; static lean_object* l_Lean_Elab_Tactic_Omega_Justification_bmodProof___closed__23; LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Tactic_Omega_Problem_fourierMotzkinData___spec__3(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Core___hyg_5____closed__8; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___lambda__1___boxed(lean_object**); LEAN_EXPORT uint8_t l_List_elem___at_Lean_Elab_Tactic_Omega_Problem_insertConstraint___spec__10(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_Omega_Problem_fourierMotzkinSelect___spec__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Core___hyg_5____closed__17; @@ -392,12 +398,13 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Justification_comboProof(lean_ static lean_object* l_Lean_Elab_Tactic_Omega_instToExprConstraint___lambda__1___closed__1; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___closed__6; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_runOmega___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_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Core___hyg_5____closed__1; static lean_object* l___auto____x40_Lean_Elab_Tactic_Omega_Core___hyg_1659____closed__4; LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Elab_Tactic_Omega_Problem_instToStringProblem___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_toString___at_Lean_Elab_Tactic_Omega_Justification_toString___spec__1(lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_Justification_bmodProof___closed__14; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_proveFalse(lean_object*, lean_object*, lean_object*, 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_Omega_Problem_proveFalse(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*); static lean_object* l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___lambda__1___closed__8; LEAN_EXPORT lean_object* l_List_replace___at_Lean_Elab_Tactic_Omega_Problem_insertConstraint___spec__15(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_Justification_bmodProof___closed__17; @@ -427,8 +434,8 @@ uint8_t lean_int_dec_lt(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Omega_Core_0__Lean_Elab_Tactic_Omega_Justification_bullet___closed__2; lean_object* lean_nat_sub(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_Justification_bmodProof___closed__6; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_solveEqualities(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_Omega_Problem_dealWithHardEquality___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_Tactic_Omega_Problem_solveEqualities(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_Lean_throwError___at_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___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_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Justification_toString(lean_object*, lean_object*, lean_object*); lean_object* lean_nat_mul(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_Problem_proveFalse___closed__4; @@ -446,7 +453,7 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Tactic_Omega_Prob lean_object* l_Lean_Omega_Constraint_combine(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_Justification_tidyProof___closed__6; lean_object* l_String_intercalate(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_addInequality__proof(lean_object*, lean_object*, 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_Omega_Problem_addInequality__proof(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_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Tactic_Omega_Problem_fourierMotzkinData___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkHashMapImp___rarg(lean_object*); static lean_object* l_List_toString___at_Lean_Elab_Tactic_Omega_Justification_toString___spec__1___closed__1; @@ -459,7 +466,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_fourierMotzkinSelect__ static lean_object* l_Lean_Elab_Tactic_Omega_instToExprConstraint___closed__3; static lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1___closed__5; static lean_object* l_Lean_Elab_Tactic_Omega_Problem_runOmega___closed__9; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_solveEquality___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_Omega_Problem_solveEquality___boxed(lean_object*, 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_Omega_Problem_explanation_x3f___default___closed__2; LEAN_EXPORT lean_object* l_Lean_HashMapImp_moveEntries___at_Lean_Elab_Tactic_Omega_Problem_insertConstraint___spec__5(lean_object*, lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); @@ -482,7 +489,7 @@ static lean_object* l_Lean_Elab_Tactic_Omega_Problem_fourierMotzkinSelect___lamb lean_object* lean_string_append(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_List_foldl___at_Lean_Elab_Tactic_Omega_Justification_toString___spec__2(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_runOmega___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_Omega_Problem_runOmega___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_EXPORT lean_object* l_Lean_HashMap_toList___at_Lean_Elab_Tactic_Omega_Problem_instToStringProblem___spec__1(lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_Problem_runOmega___closed__4; lean_object* lean_array_get_size(lean_object*); @@ -516,6 +523,7 @@ lean_object* l_List_minNatAbs(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_assumptions___default; LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Elab_Tactic_Omega_Problem_selectEquality___spec__1(lean_object*, lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Omega_Justification_proof___rarg___closed__5; static lean_object* l_Lean_Elab_Tactic_Omega_instToExprConstraint___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Lean_HashSetImp_expand___at_Lean_Elab_Tactic_Omega_Problem_insertConstraint___spec__11(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_eliminations___default; @@ -534,7 +542,7 @@ static lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Ela LEAN_EXPORT lean_object* l___auto____x40_Lean_Elab_Tactic_Omega_Core___hyg_1659_; LEAN_EXPORT lean_object* l_Lean_mkHashSet___at_Lean_Elab_Tactic_Omega_Problem_equalities___default___spec__1(lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_Justification_toString___closed__11; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_addInequality__proof___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_Omega_Problem_addInequality__proof___boxed(lean_object*, 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_Omega_Justification_bmodProof___closed__8; static lean_object* l_Lean_Elab_Tactic_Omega_Justification_toString___closed__9; lean_object* l___private_Lean_Data_HashMap_0__Lean_numBucketsForCapacity(lean_object*); @@ -5080,446 +5088,470 @@ return x_2; static lean_object* _init_l_Lean_Elab_Tactic_Omega_Justification_proof___rarg___closed__4() { _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_Elab_Tactic_Omega_Justification_proof___rarg___closed__3; +x_2 = l_ReaderT_instMonadReaderT___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_Omega_Justification_proof___rarg___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Tactic_Omega_Justification_proof___rarg___closed__4; +x_2 = l_ReaderT_instMonadReaderT___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_Omega_Justification_proof___rarg___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Tactic_Omega_Justification_proof___rarg___closed__5; x_2 = l_Lean_instInhabitedExpr; x_3 = l_instInhabited___rarg(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Justification_proof___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_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Justification_proof___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, 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) { _start: { switch (lean_obj_tag(x_4)) { case 0: { -lean_object* x_13; lean_object* x_14; uint8_t x_15; +lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_dec(x_2); -x_13 = lean_ctor_get(x_4, 2); -lean_inc(x_13); +x_15 = lean_ctor_get(x_4, 2); +lean_inc(x_15); lean_dec(x_4); -x_14 = lean_array_get_size(x_3); -x_15 = lean_nat_dec_lt(x_13, x_14); -lean_dec(x_14); -if (x_15 == 0) +x_16 = lean_array_get_size(x_3); +x_17 = lean_nat_dec_lt(x_15, x_16); +lean_dec(x_16); +if (x_17 == 0) { -lean_object* x_16; lean_object* x_17; lean_object* x_18; -lean_dec(x_13); -x_16 = l_Lean_Elab_Tactic_Omega_Justification_proof___rarg___closed__4; -x_17 = l___private_Init_Util_0__outOfBounds___rarg(x_16); -x_18 = lean_apply_8(x_17, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_18; +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +lean_dec(x_15); +x_18 = l_Lean_Elab_Tactic_Omega_Justification_proof___rarg___closed__6; +x_19 = l___private_Init_Util_0__outOfBounds___rarg(x_18); +x_20 = lean_box(x_8); +x_21 = lean_apply_10(x_19, x_5, x_6, x_7, x_20, x_9, x_10, x_11, x_12, x_13, x_14); +return x_21; } else { -lean_object* x_19; lean_object* x_20; -x_19 = lean_array_fget(x_3, x_13); -lean_dec(x_13); -x_20 = lean_apply_8(x_19, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_20; +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_array_fget(x_3, x_15); +lean_dec(x_15); +x_23 = lean_box(x_8); +x_24 = lean_apply_10(x_22, x_5, x_6, x_7, x_23, x_9, x_10, x_11, x_12, x_13, x_14); +return x_24; } } case 1: { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -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_4, 2); -lean_inc(x_23); +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_25 = lean_ctor_get(x_4, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_4, 1); +lean_inc(x_26); +x_27 = lean_ctor_get(x_4, 2); +lean_inc(x_27); lean_dec(x_4); lean_inc(x_2); -x_24 = l_Lean_Elab_Tactic_Omega_Justification_proof___rarg(x_22, x_2, x_3, x_23, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_24) == 0) +x_28 = l_Lean_Elab_Tactic_Omega_Justification_proof___rarg(x_26, x_2, x_3, x_27, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_28) == 0) { -uint8_t x_25; -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +uint8_t x_29; +x_29 = !lean_is_exclusive(x_28); +if (x_29 == 0) { -lean_object* x_26; lean_object* x_27; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_Lean_Elab_Tactic_Omega_Justification_tidyProof(x_21, x_22, x_2, x_26); -lean_dec(x_22); -lean_dec(x_21); -lean_ctor_set(x_24, 0, x_27); -return x_24; +lean_object* x_30; lean_object* x_31; +x_30 = lean_ctor_get(x_28, 0); +x_31 = l_Lean_Elab_Tactic_Omega_Justification_tidyProof(x_25, x_26, x_2, x_30); +lean_dec(x_26); +lean_dec(x_25); +lean_ctor_set(x_28, 0, x_31); +return x_28; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_28 = lean_ctor_get(x_24, 0); -x_29 = lean_ctor_get(x_24, 1); -lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_24); -x_30 = l_Lean_Elab_Tactic_Omega_Justification_tidyProof(x_21, x_22, x_2, x_28); -lean_dec(x_22); -lean_dec(x_21); -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; +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_32 = lean_ctor_get(x_28, 0); +x_33 = lean_ctor_get(x_28, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_28); +x_34 = l_Lean_Elab_Tactic_Omega_Justification_tidyProof(x_25, x_26, x_2, x_32); +lean_dec(x_26); +lean_dec(x_25); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_33); +return x_35; } } else { -uint8_t x_32; -lean_dec(x_22); -lean_dec(x_21); +uint8_t x_36; +lean_dec(x_26); +lean_dec(x_25); lean_dec(x_2); -x_32 = !lean_is_exclusive(x_24); -if (x_32 == 0) +x_36 = !lean_is_exclusive(x_28); +if (x_36 == 0) { -return x_24; +return x_28; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_24, 0); -x_34 = lean_ctor_get(x_24, 1); -lean_inc(x_34); -lean_inc(x_33); -lean_dec(x_24); -x_35 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -return x_35; +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_28, 0); +x_38 = lean_ctor_get(x_28, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_28); +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; } } } case 2: { -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_36 = lean_ctor_get(x_4, 0); -lean_inc(x_36); -x_37 = lean_ctor_get(x_4, 1); -lean_inc(x_37); -x_38 = lean_ctor_get(x_4, 3); -lean_inc(x_38); -x_39 = lean_ctor_get(x_4, 4); -lean_inc(x_39); +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_40 = lean_ctor_get(x_4, 0); +lean_inc(x_40); +x_41 = lean_ctor_get(x_4, 1); +lean_inc(x_41); +x_42 = lean_ctor_get(x_4, 3); +lean_inc(x_42); +x_43 = lean_ctor_get(x_4, 4); +lean_inc(x_43); lean_dec(x_4); +lean_inc(x_13); +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); -lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_2); -x_40 = l_Lean_Elab_Tactic_Omega_Justification_proof___rarg(x_1, x_2, x_3, x_38, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_40) == 0) +x_44 = l_Lean_Elab_Tactic_Omega_Justification_proof___rarg(x_1, x_2, x_3, x_42, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_44) == 0) { -lean_object* x_41; lean_object* x_42; lean_object* x_43; -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); +lean_object* x_45; lean_object* x_46; lean_object* x_47; +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); lean_inc(x_2); -x_43 = l_Lean_Elab_Tactic_Omega_Justification_proof___rarg(x_1, x_2, x_3, x_39, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_42); -if (lean_obj_tag(x_43) == 0) +x_47 = l_Lean_Elab_Tactic_Omega_Justification_proof___rarg(x_1, x_2, x_3, x_43, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_46); +if (lean_obj_tag(x_47) == 0) { -uint8_t x_44; -x_44 = !lean_is_exclusive(x_43); -if (x_44 == 0) +uint8_t x_48; +x_48 = !lean_is_exclusive(x_47); +if (x_48 == 0) { -lean_object* x_45; lean_object* x_46; -x_45 = lean_ctor_get(x_43, 0); -x_46 = l_Lean_Elab_Tactic_Omega_Justification_combineProof(x_36, x_37, x_1, x_2, x_41, x_45); -lean_dec(x_37); -lean_dec(x_36); -lean_ctor_set(x_43, 0, x_46); -return x_43; +lean_object* x_49; lean_object* x_50; +x_49 = lean_ctor_get(x_47, 0); +x_50 = l_Lean_Elab_Tactic_Omega_Justification_combineProof(x_40, x_41, x_1, x_2, x_45, x_49); +lean_dec(x_41); +lean_dec(x_40); +lean_ctor_set(x_47, 0, x_50); +return x_47; } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_47 = lean_ctor_get(x_43, 0); -x_48 = lean_ctor_get(x_43, 1); -lean_inc(x_48); -lean_inc(x_47); -lean_dec(x_43); -x_49 = l_Lean_Elab_Tactic_Omega_Justification_combineProof(x_36, x_37, x_1, x_2, x_41, x_47); -lean_dec(x_37); -lean_dec(x_36); -x_50 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_50, 0, x_49); -lean_ctor_set(x_50, 1, x_48); -return x_50; +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_51 = lean_ctor_get(x_47, 0); +x_52 = lean_ctor_get(x_47, 1); +lean_inc(x_52); +lean_inc(x_51); +lean_dec(x_47); +x_53 = l_Lean_Elab_Tactic_Omega_Justification_combineProof(x_40, x_41, x_1, x_2, x_45, x_51); +lean_dec(x_41); +lean_dec(x_40); +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_51; +uint8_t x_55; +lean_dec(x_45); lean_dec(x_41); -lean_dec(x_37); -lean_dec(x_36); +lean_dec(x_40); lean_dec(x_2); -x_51 = !lean_is_exclusive(x_43); -if (x_51 == 0) +x_55 = !lean_is_exclusive(x_47); +if (x_55 == 0) { -return x_43; +return x_47; } else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_ctor_get(x_43, 0); -x_53 = lean_ctor_get(x_43, 1); -lean_inc(x_53); -lean_inc(x_52); -lean_dec(x_43); -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_56; lean_object* x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_47, 0); +x_57 = lean_ctor_get(x_47, 1); +lean_inc(x_57); +lean_inc(x_56); +lean_dec(x_47); +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 { -uint8_t x_55; -lean_dec(x_39); -lean_dec(x_37); -lean_dec(x_36); +uint8_t x_59; +lean_dec(x_43); +lean_dec(x_41); +lean_dec(x_40); +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_2); -x_55 = !lean_is_exclusive(x_40); -if (x_55 == 0) +x_59 = !lean_is_exclusive(x_44); +if (x_59 == 0) { -return x_40; +return x_44; } else { -lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_56 = lean_ctor_get(x_40, 0); -x_57 = lean_ctor_get(x_40, 1); -lean_inc(x_57); -lean_inc(x_56); -lean_dec(x_40); -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; +lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_60 = lean_ctor_get(x_44, 0); +x_61 = lean_ctor_get(x_44, 1); +lean_inc(x_61); +lean_inc(x_60); +lean_dec(x_44); +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; } } } case 3: { -lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_59 = lean_ctor_get(x_4, 0); -lean_inc(x_59); -x_60 = lean_ctor_get(x_4, 1); -lean_inc(x_60); -x_61 = lean_ctor_get(x_4, 2); -lean_inc(x_61); -x_62 = lean_ctor_get(x_4, 3); -lean_inc(x_62); -x_63 = lean_ctor_get(x_4, 4); +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_63 = lean_ctor_get(x_4, 0); lean_inc(x_63); -x_64 = lean_ctor_get(x_4, 5); +x_64 = lean_ctor_get(x_4, 1); lean_inc(x_64); -x_65 = lean_ctor_get(x_4, 6); +x_65 = lean_ctor_get(x_4, 2); lean_inc(x_65); -x_66 = lean_ctor_get(x_4, 7); +x_66 = lean_ctor_get(x_4, 3); lean_inc(x_66); +x_67 = lean_ctor_get(x_4, 4); +lean_inc(x_67); +x_68 = lean_ctor_get(x_4, 5); +lean_inc(x_68); +x_69 = lean_ctor_get(x_4, 6); +lean_inc(x_69); +x_70 = lean_ctor_get(x_4, 7); +lean_inc(x_70); lean_dec(x_4); +lean_inc(x_13); +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); -lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_2); -x_67 = l_Lean_Elab_Tactic_Omega_Justification_proof___rarg(x_61, x_2, x_3, x_64, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_67) == 0) +x_71 = l_Lean_Elab_Tactic_Omega_Justification_proof___rarg(x_65, x_2, x_3, x_68, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_71) == 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_object* x_72; lean_object* x_73; lean_object* 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); lean_inc(x_2); -x_70 = l_Lean_Elab_Tactic_Omega_Justification_proof___rarg(x_62, x_2, x_3, x_66, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_69); -if (lean_obj_tag(x_70) == 0) +x_74 = l_Lean_Elab_Tactic_Omega_Justification_proof___rarg(x_66, x_2, x_3, x_70, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_73); +if (lean_obj_tag(x_74) == 0) { -uint8_t x_71; -x_71 = !lean_is_exclusive(x_70); -if (x_71 == 0) +uint8_t x_75; +x_75 = !lean_is_exclusive(x_74); +if (x_75 == 0) { -lean_object* x_72; lean_object* x_73; -x_72 = lean_ctor_get(x_70, 0); -x_73 = l_Lean_Elab_Tactic_Omega_Justification_comboProof(x_59, x_60, x_63, x_61, x_65, x_62, x_2, x_68, x_72); -lean_dec(x_62); +lean_object* x_76; lean_object* x_77; +x_76 = lean_ctor_get(x_74, 0); +x_77 = l_Lean_Elab_Tactic_Omega_Justification_comboProof(x_63, x_64, x_67, x_65, x_69, x_66, x_2, x_72, x_76); +lean_dec(x_66); +lean_dec(x_69); lean_dec(x_65); -lean_dec(x_61); +lean_dec(x_67); +lean_dec(x_64); lean_dec(x_63); -lean_dec(x_60); -lean_dec(x_59); -lean_ctor_set(x_70, 0, x_73); -return x_70; +lean_ctor_set(x_74, 0, x_77); +return x_74; } else { -lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; -x_74 = lean_ctor_get(x_70, 0); -x_75 = lean_ctor_get(x_70, 1); -lean_inc(x_75); -lean_inc(x_74); -lean_dec(x_70); -x_76 = l_Lean_Elab_Tactic_Omega_Justification_comboProof(x_59, x_60, x_63, x_61, x_65, x_62, x_2, x_68, x_74); -lean_dec(x_62); +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_78 = lean_ctor_get(x_74, 0); +x_79 = lean_ctor_get(x_74, 1); +lean_inc(x_79); +lean_inc(x_78); +lean_dec(x_74); +x_80 = l_Lean_Elab_Tactic_Omega_Justification_comboProof(x_63, x_64, x_67, x_65, x_69, x_66, x_2, x_72, x_78); +lean_dec(x_66); +lean_dec(x_69); lean_dec(x_65); -lean_dec(x_61); +lean_dec(x_67); +lean_dec(x_64); lean_dec(x_63); -lean_dec(x_60); -lean_dec(x_59); -x_77 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_77, 0, x_76); -lean_ctor_set(x_77, 1, x_75); -return x_77; +x_81 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_81, 0, x_80); +lean_ctor_set(x_81, 1, x_79); +return x_81; } } else { -uint8_t x_78; -lean_dec(x_68); +uint8_t x_82; +lean_dec(x_72); +lean_dec(x_69); +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_59); lean_dec(x_2); -x_78 = !lean_is_exclusive(x_70); -if (x_78 == 0) +x_82 = !lean_is_exclusive(x_74); +if (x_82 == 0) { -return x_70; +return x_74; } else { -lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_79 = lean_ctor_get(x_70, 0); -x_80 = lean_ctor_get(x_70, 1); -lean_inc(x_80); -lean_inc(x_79); -lean_dec(x_70); -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; +lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_83 = lean_ctor_get(x_74, 0); +x_84 = lean_ctor_get(x_74, 1); +lean_inc(x_84); +lean_inc(x_83); +lean_dec(x_74); +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_82; +uint8_t x_86; +lean_dec(x_70); +lean_dec(x_69); +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_59); +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_2); -x_82 = !lean_is_exclusive(x_67); -if (x_82 == 0) +x_86 = !lean_is_exclusive(x_71); +if (x_86 == 0) { -return x_67; +return x_71; } else { -lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_83 = lean_ctor_get(x_67, 0); -x_84 = lean_ctor_get(x_67, 1); -lean_inc(x_84); -lean_inc(x_83); -lean_dec(x_67); -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; +lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_87 = lean_ctor_get(x_71, 0); +x_88 = lean_ctor_get(x_71, 1); +lean_inc(x_88); +lean_inc(x_87); +lean_dec(x_71); +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; } } } default: { -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_86 = lean_ctor_get(x_4, 0); -lean_inc(x_86); -x_87 = lean_ctor_get(x_4, 1); -lean_inc(x_87); -x_88 = lean_ctor_get(x_4, 2); -lean_inc(x_88); -x_89 = lean_ctor_get(x_4, 3); -lean_inc(x_89); -x_90 = lean_ctor_get(x_4, 4); +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_90 = lean_ctor_get(x_4, 0); lean_inc(x_90); +x_91 = lean_ctor_get(x_4, 1); +lean_inc(x_91); +x_92 = lean_ctor_get(x_4, 2); +lean_inc(x_92); +x_93 = lean_ctor_get(x_4, 3); +lean_inc(x_93); +x_94 = lean_ctor_get(x_4, 4); +lean_inc(x_94); lean_dec(x_4); +lean_inc(x_13); +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); lean_inc(x_2); -x_91 = l_Lean_Elab_Tactic_Omega_Justification_proof___rarg(x_89, x_2, x_3, x_90, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_91) == 0) +x_95 = l_Lean_Elab_Tactic_Omega_Justification_proof___rarg(x_93, x_2, x_3, x_94, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_95) == 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_inc(x_93); +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_95, 0); +lean_inc(x_96); +x_97 = lean_ctor_get(x_95, 1); +lean_inc(x_97); +lean_dec(x_95); +x_98 = l_Lean_Elab_Tactic_Omega_Justification_bmodProof(x_90, x_91, x_92, x_93, x_2, x_96, x_10, x_11, x_12, x_13, x_97); +lean_dec(x_93); lean_dec(x_91); -x_94 = l_Lean_Elab_Tactic_Omega_Justification_bmodProof(x_86, x_87, x_88, x_89, x_2, x_92, x_8, x_9, x_10, x_11, x_93); -lean_dec(x_89); -lean_dec(x_87); -return x_94; +return x_98; } else { -uint8_t x_95; -lean_dec(x_89); -lean_dec(x_88); -lean_dec(x_87); -lean_dec(x_86); +uint8_t x_99; +lean_dec(x_93); +lean_dec(x_92); +lean_dec(x_91); +lean_dec(x_90); +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_2); -x_95 = !lean_is_exclusive(x_91); -if (x_95 == 0) +x_99 = !lean_is_exclusive(x_95); +if (x_99 == 0) { -return x_91; +return x_95; } else { -lean_object* x_96; lean_object* x_97; lean_object* x_98; -x_96 = lean_ctor_get(x_91, 0); -x_97 = lean_ctor_get(x_91, 1); -lean_inc(x_97); -lean_inc(x_96); -lean_dec(x_91); -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_100; lean_object* x_101; lean_object* x_102; +x_100 = lean_ctor_get(x_95, 0); +x_101 = lean_ctor_get(x_95, 1); +lean_inc(x_101); +lean_inc(x_100); +lean_dec(x_95); +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; } } } @@ -5530,18 +5562,20 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Justification_proof(lean_objec _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_Justification_proof___rarg___boxed), 12, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_Justification_proof___rarg___boxed), 14, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Justification_proof___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_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Justification_proof___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_object* x_14) { _start: { -lean_object* x_13; -x_13 = l_Lean_Elab_Tactic_Omega_Justification_proof___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); +uint8_t x_15; lean_object* x_16; +x_15 = lean_unbox(x_8); +lean_dec(x_8); +x_16 = l_Lean_Elab_Tactic_Omega_Justification_proof___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_15, x_9, x_10, x_11, x_12, x_13, x_14); lean_dec(x_3); lean_dec(x_1); -return x_13; +return x_16; } } LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Justification_proof___boxed(lean_object* x_1) { @@ -7173,311 +7207,314 @@ x_3 = l_Lean_Expr_app___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_proveFalse(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_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_Tactic_Omega_Problem_proveFalse(lean_object* x_1, 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) { _start: { -lean_object* x_13; +lean_object* x_15; +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_13 = l_Lean_Elab_Tactic_Omega_atomsCoeffs___rarg(x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_13) == 0) +x_15 = l_Lean_Elab_Tactic_Omega_atomsCoeffs___rarg(x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +lean_inc(x_13); +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_14); -x_16 = l_Lean_Elab_Tactic_Omega_Justification_proof___rarg(x_2, x_14, x_4, x_3, x_5, x_6, x_7, x_8, x_9, 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; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_50; -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_Elab_Tactic_Omega_Justification_bmodProof___closed__2; -x_20 = l_Lean_Elab_Tactic_Omega_Justification_bmodProof___closed__3; -x_21 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1(x_19, x_20, x_2); -x_22 = lean_ctor_get(x_1, 0); -x_23 = lean_ctor_get(x_1, 1); -if (lean_obj_tag(x_22) == 0) -{ -lean_object* x_75; -x_75 = l_Lean_Elab_Tactic_Omega_Problem_proveFalse___closed__11; -x_50 = x_75; -goto block_74; -} -else +lean_inc(x_16); +x_18 = l_Lean_Elab_Tactic_Omega_Justification_proof___rarg(x_2, x_16, x_4, x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_17); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_76; lean_object* x_77; uint8_t x_78; -x_76 = lean_ctor_get(x_22, 0); -x_77 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1___closed__1; -x_78 = lean_int_dec_le(x_77, x_76); -if (x_78 == 0) +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_52; +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_Elab_Tactic_Omega_Justification_bmodProof___closed__2; +x_22 = l_Lean_Elab_Tactic_Omega_Justification_bmodProof___closed__3; +x_23 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1(x_21, x_22, x_2); +x_24 = lean_ctor_get(x_1, 0); +x_25 = lean_ctor_get(x_1, 1); +if (lean_obj_tag(x_24) == 0) { -lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; -x_79 = lean_int_neg(x_76); -x_80 = l_Int_toNat(x_79); -lean_dec(x_79); -x_81 = l_Lean_instToExprInt_mkNat(x_80); -x_82 = l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___lambda__1___closed__16; -x_83 = l_Lean_Elab_Tactic_Omega_Justification_bmodProof___closed__1; -x_84 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1___closed__13; -x_85 = l_Lean_mkApp3(x_82, x_83, x_84, x_81); -x_86 = l_Lean_Elab_Tactic_Omega_Justification_tidyProof___closed__7; -x_87 = l_Lean_mkAppB(x_86, x_83, x_85); -x_50 = x_87; -goto block_74; +lean_object* x_77; +x_77 = l_Lean_Elab_Tactic_Omega_Problem_proveFalse___closed__11; +x_52 = x_77; +goto block_76; } else { -lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; -x_88 = l_Int_toNat(x_76); -x_89 = l_Lean_instToExprInt_mkNat(x_88); -x_90 = l_Lean_Elab_Tactic_Omega_Justification_tidyProof___closed__7; -x_91 = l_Lean_Elab_Tactic_Omega_Justification_bmodProof___closed__1; -x_92 = l_Lean_mkAppB(x_90, x_91, x_89); -x_50 = x_92; -goto block_74; +lean_object* x_78; lean_object* x_79; uint8_t x_80; +x_78 = lean_ctor_get(x_24, 0); +x_79 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1___closed__1; +x_80 = lean_int_dec_le(x_79, x_78); +if (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; +x_81 = lean_int_neg(x_78); +x_82 = l_Int_toNat(x_81); +lean_dec(x_81); +x_83 = l_Lean_instToExprInt_mkNat(x_82); +x_84 = l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___lambda__1___closed__16; +x_85 = l_Lean_Elab_Tactic_Omega_Justification_bmodProof___closed__1; +x_86 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1___closed__13; +x_87 = l_Lean_mkApp3(x_84, x_85, x_86, x_83); +x_88 = l_Lean_Elab_Tactic_Omega_Justification_tidyProof___closed__7; +x_89 = l_Lean_mkAppB(x_88, x_85, x_87); +x_52 = x_89; +goto block_76; +} +else +{ +lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_90 = l_Int_toNat(x_78); +x_91 = l_Lean_instToExprInt_mkNat(x_90); +x_92 = l_Lean_Elab_Tactic_Omega_Justification_tidyProof___closed__7; +x_93 = l_Lean_Elab_Tactic_Omega_Justification_bmodProof___closed__1; +x_94 = l_Lean_mkAppB(x_92, x_93, x_91); +x_52 = x_94; +goto block_76; } } -block_49: +block_51: { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_25 = l_Lean_Elab_Tactic_Omega_Problem_proveFalse___closed__3; -lean_inc(x_24); -x_26 = l_Lean_Expr_app___override(x_25, x_24); -x_27 = l_Lean_Elab_Tactic_Omega_Problem_proveFalse___closed__7; +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_27 = l_Lean_Elab_Tactic_Omega_Problem_proveFalse___closed__3; +lean_inc(x_26); +x_28 = l_Lean_Expr_app___override(x_27, x_26); +x_29 = l_Lean_Elab_Tactic_Omega_Problem_proveFalse___closed__7; +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_28 = l_Lean_Meta_mkEq(x_26, x_27, x_8, x_9, x_10, x_11, x_18); -if (lean_obj_tag(x_28) == 0) +x_30 = l_Lean_Meta_mkEq(x_28, x_29, x_10, x_11, x_12, x_13, x_20); +if (lean_obj_tag(x_30) == 0) { -lean_object* x_29; lean_object* x_30; lean_object* x_31; -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_Lean_Meta_mkDecideProof(x_29, x_8, x_9, x_10, x_11, x_30); -if (lean_obj_tag(x_31) == 0) +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +x_33 = l_Lean_Meta_mkDecideProof(x_31, x_10, x_11, x_12, x_13, x_32); +if (lean_obj_tag(x_33) == 0) { -uint8_t x_32; -x_32 = !lean_is_exclusive(x_31); -if (x_32 == 0) +uint8_t x_34; +x_34 = !lean_is_exclusive(x_33); +if (x_34 == 0) { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_31, 0); -x_34 = l_Lean_Elab_Tactic_Omega_Problem_proveFalse___closed__10; -x_35 = l_Lean_mkApp5(x_34, x_24, x_33, x_21, x_14, x_17); -lean_ctor_set(x_31, 0, x_35); -return x_31; +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_33, 0); +x_36 = l_Lean_Elab_Tactic_Omega_Problem_proveFalse___closed__10; +x_37 = l_Lean_mkApp5(x_36, x_26, x_35, x_23, x_16, x_19); +lean_ctor_set(x_33, 0, x_37); +return x_33; } else { -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_36 = lean_ctor_get(x_31, 0); -x_37 = lean_ctor_get(x_31, 1); -lean_inc(x_37); -lean_inc(x_36); -lean_dec(x_31); -x_38 = l_Lean_Elab_Tactic_Omega_Problem_proveFalse___closed__10; -x_39 = l_Lean_mkApp5(x_38, x_24, x_36, x_21, x_14, x_17); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_39); -lean_ctor_set(x_40, 1, x_37); -return x_40; +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_38 = lean_ctor_get(x_33, 0); +x_39 = lean_ctor_get(x_33, 1); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_33); +x_40 = l_Lean_Elab_Tactic_Omega_Problem_proveFalse___closed__10; +x_41 = l_Lean_mkApp5(x_40, x_26, x_38, x_23, x_16, x_19); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_39); +return x_42; } } else { -uint8_t x_41; -lean_dec(x_24); -lean_dec(x_21); -lean_dec(x_17); -lean_dec(x_14); -x_41 = !lean_is_exclusive(x_31); -if (x_41 == 0) +uint8_t x_43; +lean_dec(x_26); +lean_dec(x_23); +lean_dec(x_19); +lean_dec(x_16); +x_43 = !lean_is_exclusive(x_33); +if (x_43 == 0) { -return x_31; +return x_33; } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_31, 0); -x_43 = lean_ctor_get(x_31, 1); -lean_inc(x_43); -lean_inc(x_42); -lean_dec(x_31); -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_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_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; } } } else { -uint8_t x_45; -lean_dec(x_24); -lean_dec(x_21); -lean_dec(x_17); -lean_dec(x_14); +uint8_t x_47; +lean_dec(x_26); +lean_dec(x_23); +lean_dec(x_19); +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); -x_45 = !lean_is_exclusive(x_28); -if (x_45 == 0) +x_47 = !lean_is_exclusive(x_30); +if (x_47 == 0) { -return x_28; +return x_30; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_46 = lean_ctor_get(x_28, 0); -x_47 = lean_ctor_get(x_28, 1); -lean_inc(x_47); -lean_inc(x_46); -lean_dec(x_28); -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_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_30, 0); +x_49 = lean_ctor_get(x_30, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_30); +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; } } } -block_74: -{ -lean_object* x_51; lean_object* x_52; -x_51 = l_Lean_Elab_Tactic_Omega_Justification_tidyProof___closed__4; -x_52 = l_Lean_Expr_app___override(x_51, x_50); -if (lean_obj_tag(x_23) == 0) +block_76: { lean_object* x_53; lean_object* x_54; -x_53 = l_Lean_Elab_Tactic_Omega_Problem_proveFalse___closed__11; -x_54 = l_Lean_Expr_app___override(x_52, x_53); -x_24 = x_54; -goto block_49; -} -else -{ -lean_object* x_55; lean_object* x_56; uint8_t x_57; -x_55 = lean_ctor_get(x_23, 0); -x_56 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1___closed__1; -x_57 = lean_int_dec_le(x_56, x_55); -if (x_57 == 0) +x_53 = l_Lean_Elab_Tactic_Omega_Justification_tidyProof___closed__4; +x_54 = l_Lean_Expr_app___override(x_53, x_52); +if (lean_obj_tag(x_25) == 0) { -lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_58 = lean_int_neg(x_55); -x_59 = l_Int_toNat(x_58); -lean_dec(x_58); -x_60 = l_Lean_instToExprInt_mkNat(x_59); -x_61 = l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___lambda__1___closed__16; -x_62 = l_Lean_Elab_Tactic_Omega_Justification_bmodProof___closed__1; -x_63 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1___closed__13; -x_64 = l_Lean_mkApp3(x_61, x_62, x_63, x_60); -x_65 = l_Lean_Elab_Tactic_Omega_Justification_tidyProof___closed__7; -x_66 = l_Lean_mkAppB(x_65, x_62, x_64); -x_67 = l_Lean_Expr_app___override(x_52, x_66); -x_24 = x_67; -goto block_49; +lean_object* x_55; lean_object* x_56; +x_55 = l_Lean_Elab_Tactic_Omega_Problem_proveFalse___closed__11; +x_56 = l_Lean_Expr_app___override(x_54, x_55); +x_26 = x_56; +goto block_51; } else { -lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; -x_68 = l_Int_toNat(x_55); -x_69 = l_Lean_instToExprInt_mkNat(x_68); -x_70 = l_Lean_Elab_Tactic_Omega_Justification_tidyProof___closed__7; -x_71 = l_Lean_Elab_Tactic_Omega_Justification_bmodProof___closed__1; -x_72 = l_Lean_mkAppB(x_70, x_71, x_69); -x_73 = l_Lean_Expr_app___override(x_52, x_72); -x_24 = x_73; -goto block_49; +lean_object* x_57; lean_object* x_58; uint8_t x_59; +x_57 = lean_ctor_get(x_25, 0); +x_58 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1___closed__1; +x_59 = lean_int_dec_le(x_58, x_57); +if (x_59 == 0) +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_60 = lean_int_neg(x_57); +x_61 = l_Int_toNat(x_60); +lean_dec(x_60); +x_62 = l_Lean_instToExprInt_mkNat(x_61); +x_63 = l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___lambda__1___closed__16; +x_64 = l_Lean_Elab_Tactic_Omega_Justification_bmodProof___closed__1; +x_65 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1___closed__13; +x_66 = l_Lean_mkApp3(x_63, x_64, x_65, x_62); +x_67 = l_Lean_Elab_Tactic_Omega_Justification_tidyProof___closed__7; +x_68 = l_Lean_mkAppB(x_67, x_64, x_66); +x_69 = l_Lean_Expr_app___override(x_54, x_68); +x_26 = x_69; +goto block_51; +} +else +{ +lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_70 = l_Int_toNat(x_57); +x_71 = l_Lean_instToExprInt_mkNat(x_70); +x_72 = l_Lean_Elab_Tactic_Omega_Justification_tidyProof___closed__7; +x_73 = l_Lean_Elab_Tactic_Omega_Justification_bmodProof___closed__1; +x_74 = l_Lean_mkAppB(x_72, x_73, x_71); +x_75 = l_Lean_Expr_app___override(x_54, x_74); +x_26 = x_75; +goto block_51; } } } } else { -uint8_t x_93; -lean_dec(x_14); +uint8_t x_95; +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); -x_93 = !lean_is_exclusive(x_16); -if (x_93 == 0) +x_95 = !lean_is_exclusive(x_18); +if (x_95 == 0) { -return x_16; +return x_18; } else { -lean_object* x_94; lean_object* x_95; lean_object* x_96; -x_94 = lean_ctor_get(x_16, 0); -x_95 = lean_ctor_get(x_16, 1); -lean_inc(x_95); -lean_inc(x_94); -lean_dec(x_16); -x_96 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_96, 0, x_94); -lean_ctor_set(x_96, 1, x_95); -return x_96; +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_18, 0); +x_97 = lean_ctor_get(x_18, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_18); +x_98 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +return x_98; } } } else { -uint8_t x_97; +uint8_t x_99; +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_97 = !lean_is_exclusive(x_13); -if (x_97 == 0) +x_99 = !lean_is_exclusive(x_15); +if (x_99 == 0) { -return x_13; +return x_15; } else { -lean_object* x_98; lean_object* x_99; lean_object* x_100; -x_98 = lean_ctor_get(x_13, 0); -x_99 = lean_ctor_get(x_13, 1); -lean_inc(x_99); -lean_inc(x_98); -lean_dec(x_13); -x_100 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_100, 0, x_98); -lean_ctor_set(x_100, 1, x_99); -return x_100; +lean_object* x_100; lean_object* x_101; lean_object* x_102; +x_100 = lean_ctor_get(x_15, 0); +x_101 = lean_ctor_get(x_15, 1); +lean_inc(x_101); +lean_inc(x_100); +lean_dec(x_15); +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_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_proveFalse___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_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_Tactic_Omega_Problem_proveFalse___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -lean_object* x_13; -x_13 = l_Lean_Elab_Tactic_Omega_Problem_proveFalse(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +uint8_t x_15; lean_object* x_16; +x_15 = lean_unbox(x_8); +lean_dec(x_8); +x_16 = l_Lean_Elab_Tactic_Omega_Problem_proveFalse(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_15, x_9, x_10, x_11, x_12, x_13, x_14); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -return x_13; +return x_16; } } LEAN_EXPORT uint64_t l_List_foldl___at_Lean_Elab_Tactic_Omega_Problem_insertConstraint___spec__2(uint64_t x_1, lean_object* x_2) { @@ -8464,7 +8501,7 @@ lean_inc(x_38); lean_inc(x_5); lean_inc(x_3); lean_inc(x_4); -x_41 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_Problem_proveFalse___boxed), 12, 4); +x_41 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_Problem_proveFalse___boxed), 14, 4); lean_closure_set(x_41, 0, x_4); lean_closure_set(x_41, 1, x_3); lean_closure_set(x_41, 2, x_5); @@ -8500,7 +8537,7 @@ lean_inc(x_46); lean_inc(x_5); lean_inc(x_3); lean_inc(x_4); -x_51 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_Problem_proveFalse___boxed), 12, 4); +x_51 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_Problem_proveFalse___boxed), 14, 4); lean_closure_set(x_51, 0, x_4); lean_closure_set(x_51, 1, x_3); lean_closure_set(x_51, 2, x_5); @@ -9760,107 +9797,107 @@ lean_dec(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___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_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___spec__1(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) { _start: { -lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = lean_ctor_get(x_7, 5); -x_11 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_5, x_6, x_7, x_8, x_9); -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) +lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_12 = lean_ctor_get(x_9, 5); +x_13 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_7, x_8, x_9, x_10, x_11); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) { -lean_object* x_13; lean_object* x_14; -x_13 = lean_ctor_get(x_11, 0); -lean_inc(x_10); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_10); -lean_ctor_set(x_14, 1, x_13); -lean_ctor_set_tag(x_11, 1); -lean_ctor_set(x_11, 0, x_14); -return x_11; +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_13, 0); +lean_inc(x_12); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_12); +lean_ctor_set(x_16, 1, x_15); +lean_ctor_set_tag(x_13, 1); +lean_ctor_set(x_13, 0, x_16); +return x_13; } else { -lean_object* x_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); -lean_inc(x_10); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_10); -lean_ctor_set(x_17, 1, x_15); -x_18 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_16); -return x_18; +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_17 = lean_ctor_get(x_13, 0); +x_18 = lean_ctor_get(x_13, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_13); +lean_inc(x_12); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_12); +lean_ctor_set(x_19, 1, x_17); +x_20 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_18); +return x_20; } } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___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_throwError___at_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___spec__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, 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; uint8_t x_12; -x_10 = lean_ctor_get(x_7, 5); -x_11 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_5, x_6, x_7, x_8, x_9); -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) +lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_12 = lean_ctor_get(x_9, 5); +x_13 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_7, x_8, x_9, x_10, x_11); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) { -lean_object* x_13; lean_object* x_14; -x_13 = lean_ctor_get(x_11, 0); -lean_inc(x_10); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_10); -lean_ctor_set(x_14, 1, x_13); -lean_ctor_set_tag(x_11, 1); -lean_ctor_set(x_11, 0, x_14); -return x_11; +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_13, 0); +lean_inc(x_12); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_12); +lean_ctor_set(x_16, 1, x_15); +lean_ctor_set_tag(x_13, 1); +lean_ctor_set(x_13, 0, x_16); +return x_13; } else { -lean_object* x_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); -lean_inc(x_10); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_10); -lean_ctor_set(x_17, 1, x_15); -x_18 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_16); -return x_18; +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_17 = lean_ctor_get(x_13, 0); +x_18 = lean_ctor_get(x_13, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_13); +lean_inc(x_12); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_12); +lean_ctor_set(x_19, 1, x_17); +x_20 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_18); +return x_20; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___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, 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) { _start: { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +lean_object* x_18; lean_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_3); lean_inc(x_1); -x_16 = l_Lean_Omega_bmod__coeffs(x_1, x_2, x_3); +x_18 = l_Lean_Omega_bmod__coeffs(x_1, x_2, x_3); lean_inc(x_1); -x_17 = l_Int_bmod(x_4, x_1); -x_18 = l_Lean_Omega_Constraint_exact(x_17); -x_19 = lean_alloc_ctor(4, 5, 0); -lean_ctor_set(x_19, 0, x_1); -lean_ctor_set(x_19, 1, x_4); -lean_ctor_set(x_19, 2, x_2); -lean_ctor_set(x_19, 3, x_3); -lean_ctor_set(x_19, 4, x_5); -x_20 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_20, 0, x_16); -lean_ctor_set(x_20, 1, x_18); -lean_ctor_set(x_20, 2, x_19); -x_21 = l_Lean_Elab_Tactic_Omega_Problem_addConstraint(x_6, 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_15); -return x_22; +x_19 = l_Int_bmod(x_4, x_1); +x_20 = l_Lean_Omega_Constraint_exact(x_19); +x_21 = lean_alloc_ctor(4, 5, 0); +lean_ctor_set(x_21, 0, x_1); +lean_ctor_set(x_21, 1, x_4); +lean_ctor_set(x_21, 2, x_2); +lean_ctor_set(x_21, 3, x_3); +lean_ctor_set(x_21, 4, x_5); +x_22 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_22, 0, x_18); +lean_ctor_set(x_22, 1, x_20); +lean_ctor_set(x_22, 2, x_21); +x_23 = l_Lean_Elab_Tactic_Omega_Problem_addConstraint(x_6, 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_17); +return x_24; } } static lean_object* _init_l_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___closed__1() { @@ -9924,341 +9961,351 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_Tactic_Omega_Problem_dealWithHardEquality(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_object* x_11, lean_object* x_12) { _start: { -lean_object* x_11; lean_object* x_12; -x_11 = lean_ctor_get(x_1, 2); -lean_inc(x_11); +lean_object* x_13; lean_object* x_14; +x_13 = lean_ctor_get(x_1, 2); +lean_inc(x_13); lean_inc(x_2); -x_12 = l_Lean_HashMapImp_find_x3f___at_Lean_Elab_Tactic_Omega_Problem_addConstraint___spec__1(x_11, x_2); -if (lean_obj_tag(x_12) == 0) +x_14 = l_Lean_HashMapImp_find_x3f___at_Lean_Elab_Tactic_Omega_Problem_addConstraint___spec__1(x_13, x_2); +if (lean_obj_tag(x_14) == 0) { -lean_object* x_13; +lean_object* 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_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_1); -lean_ctor_set(x_13, 1, x_10); -return x_13; +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_1); +lean_ctor_set(x_15, 1, x_12); +return x_15; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_12, 0); -lean_inc(x_14); -lean_dec(x_12); -x_15 = lean_ctor_get(x_14, 1); -lean_inc(x_15); -x_16 = lean_ctor_get(x_15, 0); +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_14, 0); lean_inc(x_16); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; -lean_dec(x_15); lean_dec(x_14); +x_17 = lean_ctor_get(x_16, 1); +lean_inc(x_17); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_19; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_1); -lean_ctor_set(x_17, 1, x_10); -return x_17; +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_1); +lean_ctor_set(x_19, 1, x_12); +return x_19; } -else -{ -lean_object* x_18; -x_18 = lean_ctor_get(x_15, 1); -lean_inc(x_18); -lean_dec(x_15); -if (lean_obj_tag(x_18) == 0) +else { -lean_object* x_19; +lean_object* x_20; +x_20 = lean_ctor_get(x_17, 1); +lean_inc(x_20); +lean_dec(x_17); +if (lean_obj_tag(x_20) == 0) +{ +lean_object* x_21; +lean_dec(x_18); lean_dec(x_16); -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_4); lean_dec(x_3); lean_dec(x_2); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_1); -lean_ctor_set(x_19, 1, x_10); -return x_19; +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_1); +lean_ctor_set(x_21, 1, x_12); +return x_21; } 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; -x_20 = lean_ctor_get(x_14, 0); -lean_inc(x_20); -x_21 = lean_ctor_get(x_14, 2); -lean_inc(x_21); -lean_dec(x_14); +lean_object* x_22; lean_object* x_23; 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_16, 0); lean_inc(x_22); -lean_dec(x_16); -x_23 = lean_ctor_get(x_18, 0); +x_23 = lean_ctor_get(x_16, 2); lean_inc(x_23); +lean_dec(x_16); +x_24 = lean_ctor_get(x_18, 0); +lean_inc(x_24); lean_dec(x_18); +x_25 = lean_ctor_get(x_20, 0); +lean_inc(x_25); +lean_dec(x_20); lean_inc(x_2); -x_24 = l_List_minNatAbs(x_2); -x_25 = lean_unsigned_to_nat(1u); -x_26 = lean_nat_add(x_24, x_25); -lean_dec(x_24); +x_26 = l_List_minNatAbs(x_2); +x_27 = lean_unsigned_to_nat(1u); +x_28 = lean_nat_add(x_26, x_27); +lean_dec(x_26); +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_27 = l_Lean_Elab_Tactic_Omega_atomsCoeffs___rarg(x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_27) == 0) +x_29 = l_Lean_Elab_Tactic_Omega_atomsCoeffs___rarg(x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_29) == 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; -x_28 = lean_ctor_get(x_27, 0); +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_29, 1); +lean_inc(x_31); +lean_dec(x_29); lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); -lean_inc(x_29); -lean_dec(x_27); -lean_inc(x_26); -x_30 = l_Lean_mkNatLit(x_26); -x_31 = l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___lambda__1___closed__10; -x_32 = l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___lambda__1___closed__14; -x_33 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1(x_31, x_32, x_2); +x_32 = l_Lean_mkNatLit(x_28); +x_33 = l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___lambda__1___closed__10; +x_34 = l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___lambda__1___closed__14; +x_35 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1(x_33, x_34, x_2); lean_dec(x_2); -x_34 = l_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___closed__1; -x_35 = l_Lean_mkApp3(x_34, x_30, x_33, x_28); +x_36 = l_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___closed__1; +x_37 = l_Lean_mkApp3(x_36, x_32, x_35, x_30); +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); lean_inc(x_3); -x_36 = l_Lean_Elab_Tactic_Omega_lookup(x_35, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_29); -if (lean_obj_tag(x_36) == 0) +x_38 = l_Lean_Elab_Tactic_Omega_lookup(x_37, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_31); +if (lean_obj_tag(x_38) == 0) { -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; -x_37 = lean_ctor_get(x_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_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_38, 0); lean_inc(x_39); -x_40 = lean_ctor_get(x_37, 1); +x_40 = lean_ctor_get(x_38, 1); lean_inc(x_40); -lean_dec(x_37); -x_41 = lean_int_dec_eq(x_23, x_22); -lean_dec(x_23); -if (x_41 == 0) -{ -lean_object* x_42; lean_object* x_43; -lean_dec(x_40); +lean_dec(x_38); +x_41 = lean_ctor_get(x_39, 0); +lean_inc(x_41); +x_42 = lean_ctor_get(x_39, 1); +lean_inc(x_42); lean_dec(x_39); -lean_dec(x_26); +x_43 = lean_int_dec_eq(x_25, x_24); +lean_dec(x_25); +if (x_43 == 0) +{ +lean_object* x_44; lean_object* x_45; +lean_dec(x_42); +lean_dec(x_41); +lean_dec(x_28); +lean_dec(x_24); +lean_dec(x_23); lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_20); lean_dec(x_1); -x_42 = l_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___closed__3; -x_43 = l_Lean_throwError___at_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___spec__1(x_42, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_38); +x_44 = l_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___closed__3; +x_45 = l_Lean_throwError___at_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___spec__1(x_44, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_40); +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); -return x_43; +return x_45; } else { -if (lean_obj_tag(x_40) == 0) +if (lean_obj_tag(x_42) == 0) { -lean_object* x_44; lean_object* x_45; uint8_t x_46; -lean_dec(x_39); -lean_dec(x_26); +lean_object* x_46; lean_object* x_47; uint8_t x_48; +lean_dec(x_41); +lean_dec(x_28); +lean_dec(x_24); +lean_dec(x_23); lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_20); lean_dec(x_1); -x_44 = l_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___closed__5; -x_45 = l_Lean_throwError___at_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___spec__2(x_44, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_38); +x_46 = l_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___closed__5; +x_47 = l_Lean_throwError___at_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___spec__2(x_46, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_40); +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_46 = !lean_is_exclusive(x_45); -if (x_46 == 0) +x_48 = !lean_is_exclusive(x_47); +if (x_48 == 0) { -return x_45; +return x_47; } 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; +lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_47, 0); +x_50 = lean_ctor_get(x_47, 1); +lean_inc(x_50); +lean_inc(x_49); +lean_dec(x_47); +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_50; lean_object* x_51; lean_object* x_52; uint8_t x_53; -x_50 = lean_ctor_get(x_40, 0); -lean_inc(x_50); -lean_dec(x_40); -x_51 = lean_ctor_get(x_50, 0); -lean_inc(x_51); -lean_dec(x_50); -x_52 = lean_unsigned_to_nat(0u); -x_53 = lean_nat_dec_eq(x_51, x_52); -lean_dec(x_51); -if (x_53 == 0) +lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; +x_52 = lean_ctor_get(x_42, 0); +lean_inc(x_52); +lean_dec(x_42); +x_53 = lean_ctor_get(x_52, 0); +lean_inc(x_53); +lean_dec(x_52); +x_54 = lean_unsigned_to_nat(0u); +x_55 = lean_nat_dec_eq(x_53, x_54); +lean_dec(x_53); +if (x_55 == 0) { -lean_object* x_54; lean_object* x_55; uint8_t x_56; -lean_dec(x_39); -lean_dec(x_26); +lean_object* x_56; lean_object* x_57; uint8_t x_58; +lean_dec(x_41); +lean_dec(x_28); +lean_dec(x_24); +lean_dec(x_23); lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_20); lean_dec(x_1); -x_54 = l_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___closed__7; -x_55 = l_Lean_throwError___at_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___spec__2(x_54, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_38); +x_56 = l_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___closed__7; +x_57 = l_Lean_throwError___at_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___spec__2(x_56, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_40); +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_56 = !lean_is_exclusive(x_55); -if (x_56 == 0) +x_58 = !lean_is_exclusive(x_57); +if (x_58 == 0) { -return x_55; +return x_57; } else { -lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_57 = lean_ctor_get(x_55, 0); -x_58 = lean_ctor_get(x_55, 1); -lean_inc(x_58); -lean_inc(x_57); -lean_dec(x_55); -x_59 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_59, 0, x_57); -lean_ctor_set(x_59, 1, x_58); -return x_59; +lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_59 = lean_ctor_get(x_57, 0); +x_60 = lean_ctor_get(x_57, 1); +lean_inc(x_60); +lean_inc(x_59); +lean_dec(x_57); +x_61 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_61, 0, x_59); +lean_ctor_set(x_61, 1, x_60); +return x_61; } } else { -lean_object* x_60; lean_object* x_61; -x_60 = lean_box(0); -x_61 = l_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___lambda__1(x_26, x_39, x_20, x_22, x_21, x_1, x_60, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_38); +lean_object* x_62; lean_object* x_63; +x_62 = lean_box(0); +x_63 = l_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___lambda__1(x_28, x_41, x_22, x_24, x_23, x_1, x_62, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_40); +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); -return x_61; +return x_63; } } } } else { -uint8_t x_62; -lean_dec(x_26); +uint8_t x_64; +lean_dec(x_28); +lean_dec(x_25); +lean_dec(x_24); lean_dec(x_23); lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_20); +lean_dec(x_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_1); -x_62 = !lean_is_exclusive(x_36); -if (x_62 == 0) +x_64 = !lean_is_exclusive(x_38); +if (x_64 == 0) { -return x_36; +return x_38; } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_63 = lean_ctor_get(x_36, 0); -x_64 = lean_ctor_get(x_36, 1); -lean_inc(x_64); -lean_inc(x_63); -lean_dec(x_36); -x_65 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_65, 0, x_63); -lean_ctor_set(x_65, 1, x_64); -return x_65; +lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_65 = lean_ctor_get(x_38, 0); +x_66 = lean_ctor_get(x_38, 1); +lean_inc(x_66); +lean_inc(x_65); +lean_dec(x_38); +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 { -uint8_t x_66; -lean_dec(x_26); +uint8_t x_68; +lean_dec(x_28); +lean_dec(x_25); +lean_dec(x_24); lean_dec(x_23); lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_20); +lean_dec(x_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_66 = !lean_is_exclusive(x_27); -if (x_66 == 0) +x_68 = !lean_is_exclusive(x_29); +if (x_68 == 0) { -return x_27; +return x_29; } else { -lean_object* x_67; lean_object* x_68; lean_object* x_69; -x_67 = lean_ctor_get(x_27, 0); -x_68 = lean_ctor_get(x_27, 1); -lean_inc(x_68); -lean_inc(x_67); -lean_dec(x_27); -x_69 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_69, 0, x_67); -lean_ctor_set(x_69, 1, x_68); -return x_69; +lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_69 = lean_ctor_get(x_29, 0); +x_70 = lean_ctor_get(x_29, 1); +lean_inc(x_70); +lean_inc(x_69); +lean_dec(x_29); +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; } } } @@ -10266,195 +10313,248 @@ return x_69; } } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___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_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___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_10; -x_10 = l_Lean_throwError___at_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +uint8_t x_12; lean_object* x_13; +x_12 = lean_unbox(x_5); +lean_dec(x_5); +x_13 = l_Lean_throwError___at_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___spec__1(x_1, x_2, x_3, x_4, x_12, 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_10; +return x_13; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___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_Lean_throwError___at_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___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) { _start: { -lean_object* x_10; -x_10 = l_Lean_throwError___at_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +uint8_t x_12; lean_object* x_13; +x_12 = lean_unbox(x_5); +lean_dec(x_5); +x_13 = l_Lean_throwError___at_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___spec__2(x_1, x_2, x_3, x_4, x_12, 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_10; +return x_13; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { -_start: -{ -lean_object* x_16; -x_16 = l_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___lambda__1___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +_start: +{ +uint8_t x_18; lean_object* x_19; +x_18 = lean_unbox(x_11); +lean_dec(x_11); +x_19 = l_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_18, x_12, x_13, x_14, x_15, x_16, x_17); +lean_dec(x_16); +lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); -lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -return x_16; +return x_19; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_solveEquality(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_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_Omega_Problem_dealWithHardEquality___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_12; uint8_t x_13; -x_12 = lean_unsigned_to_nat(1u); -x_13 = lean_nat_dec_eq(x_3, x_12); -if (x_13 == 0) -{ -lean_object* x_14; -x_14 = l_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +uint8_t x_13; lean_object* x_14; +x_13 = lean_unbox(x_6); +lean_dec(x_6); +x_14 = l_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality(x_1, x_2, x_3, x_4, x_5, x_13, x_7, x_8, x_9, x_10, x_11, x_12); return x_14; } +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_solveEquality(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_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; uint8_t x_15; +x_14 = lean_unsigned_to_nat(1u); +x_15 = lean_nat_dec_eq(x_3, x_14); +if (x_15 == 0) +{ +lean_object* x_16; +x_16 = l_Lean_Elab_Tactic_Omega_Problem_dealWithHardEquality(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_16; +} else { -lean_object* x_15; lean_object* x_16; +lean_object* x_17; lean_object* x_18; +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); -x_15 = l_Lean_Elab_Tactic_Omega_Problem_solveEasyEquality(x_1, x_2); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_11); -return x_16; +x_17 = l_Lean_Elab_Tactic_Omega_Problem_solveEasyEquality(x_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_13); +return x_18; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_solveEquality___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_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_Omega_Problem_solveEquality___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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_12; -x_12 = l_Lean_Elab_Tactic_Omega_Problem_solveEquality(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +uint8_t x_14; lean_object* x_15; +x_14 = lean_unbox(x_7); +lean_dec(x_7); +x_15 = l_Lean_Elab_Tactic_Omega_Problem_solveEquality(x_1, x_2, x_3, x_4, x_5, x_6, x_14, x_8, x_9, x_10, x_11, x_12, x_13); lean_dec(x_3); -return x_12; +return x_15; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_solveEqualities(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_solveEqualities(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) { _start: { -uint8_t x_10; -x_10 = lean_ctor_get_uint8(x_1, sizeof(void*)*7); -if (x_10 == 0) +uint8_t x_12; +x_12 = lean_ctor_get_uint8(x_1, sizeof(void*)*7); +if (x_12 == 0) { -lean_object* x_11; +lean_object* x_13; +lean_dec(x_10); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_11 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_11, 0, x_1); -lean_ctor_set(x_11, 1, x_9); -return x_11; +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_1); +lean_ctor_set(x_13, 1, x_11); +return x_13; } else { -lean_object* x_12; +lean_object* x_14; lean_inc(x_1); -x_12 = l_Lean_Elab_Tactic_Omega_Problem_selectEquality(x_1); -if (lean_obj_tag(x_12) == 0) +x_14 = l_Lean_Elab_Tactic_Omega_Problem_selectEquality(x_1); +if (lean_obj_tag(x_14) == 0) { -lean_object* x_13; +lean_object* x_15; +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_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_1); -lean_ctor_set(x_13, 1, x_9); -return x_13; +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_1); +lean_ctor_set(x_15, 1, x_11); +return x_15; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_14 = lean_ctor_get(x_12, 0); -lean_inc(x_14); -lean_dec(x_12); -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_14, 1); +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_16 = lean_ctor_get(x_14, 0); lean_inc(x_16); lean_dec(x_14); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +lean_inc(x_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_2); -x_17 = l_Lean_Elab_Tactic_Omega_Problem_solveEquality(x_1, x_15, x_16, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_16); -if (lean_obj_tag(x_17) == 0) +x_19 = l_Lean_Elab_Tactic_Omega_Problem_solveEquality(x_1, x_17, x_18, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_18); +if (lean_obj_tag(x_19) == 0) { -lean_object* x_18; lean_object* x_19; -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_1 = x_18; -x_9 = x_19; +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); +x_1 = x_20; +x_11 = x_21; goto _start; } else { -uint8_t x_21; +uint8_t x_23; +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_21 = !lean_is_exclusive(x_17); -if (x_21 == 0) +x_23 = !lean_is_exclusive(x_19); +if (x_23 == 0) { -return x_17; +return x_19; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_17, 0); -x_23 = lean_ctor_get(x_17, 1); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_17); -x_24 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_24, 0, x_22); -lean_ctor_set(x_24, 1, x_23); -return x_24; +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_19, 0); +x_25 = lean_ctor_get(x_19, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_19); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +return x_26; +} } } } } } +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_solveEqualities___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_5); +lean_dec(x_5); +x_13 = l_Lean_Elab_Tactic_Omega_Problem_solveEqualities(x_1, x_2, x_3, x_4, x_12, x_6, x_7, x_8, x_9, x_10, x_11); +return x_13; +} } static lean_object* _init_l_Lean_Elab_Tactic_Omega_Problem_addInequality__proof___closed__1() { _start: @@ -10486,173 +10586,177 @@ x_3 = l_Lean_Expr_const___override(x_2, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_addInequality__proof(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_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_Omega_Problem_addInequality__proof(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_object* x_12, lean_object* x_13) { _start: { -lean_object* x_12; +lean_object* 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); -x_12 = l_Lean_Elab_Tactic_Omega_atomsCoeffs___rarg(x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_12) == 0) -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; -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_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); -if (lean_obj_tag(x_15) == 0) +x_14 = l_Lean_Elab_Tactic_Omega_atomsCoeffs___rarg(x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_14) == 0) { -uint8_t x_16; -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* 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 = lean_box(x_7); +x_18 = lean_apply_10(x_3, x_4, x_5, x_6, x_17, x_8, x_9, x_10, x_11, x_12, x_16); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_17; lean_object* x_18; uint8_t x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_17 = lean_ctor_get(x_15, 0); -x_18 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1___closed__1; -x_19 = lean_int_dec_le(x_18, x_1); -x_20 = l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___lambda__1___closed__10; -x_21 = l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___lambda__1___closed__14; -x_22 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1(x_20, x_21, x_2); +uint8_t x_19; +x_19 = !lean_is_exclusive(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; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_23 = lean_int_neg(x_1); -x_24 = l_Int_toNat(x_23); -lean_dec(x_23); -x_25 = l_Lean_instToExprInt_mkNat(x_24); -x_26 = l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___lambda__1___closed__16; -x_27 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1___closed__10; -x_28 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1___closed__13; -x_29 = l_Lean_mkApp3(x_26, x_27, x_28, x_25); -x_30 = l_Lean_Elab_Tactic_Omega_Problem_addInequality__proof___closed__3; -x_31 = l_Lean_mkApp4(x_30, x_29, x_22, x_13, x_17); -lean_ctor_set(x_15, 0, x_31); -return x_15; +lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_20 = lean_ctor_get(x_18, 0); +x_21 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1___closed__1; +x_22 = lean_int_dec_le(x_21, x_1); +x_23 = l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___lambda__1___closed__10; +x_24 = l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___lambda__1___closed__14; +x_25 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1(x_23, x_24, x_2); +if (x_22 == 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 = lean_int_neg(x_1); +x_27 = l_Int_toNat(x_26); +lean_dec(x_26); +x_28 = l_Lean_instToExprInt_mkNat(x_27); +x_29 = l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___lambda__1___closed__16; +x_30 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1___closed__10; +x_31 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1___closed__13; +x_32 = l_Lean_mkApp3(x_29, x_30, x_31, x_28); +x_33 = l_Lean_Elab_Tactic_Omega_Problem_addInequality__proof___closed__3; +x_34 = l_Lean_mkApp4(x_33, x_32, x_25, x_15, x_20); +lean_ctor_set(x_18, 0, x_34); +return x_18; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_32 = l_Int_toNat(x_1); -x_33 = l_Lean_instToExprInt_mkNat(x_32); -x_34 = l_Lean_Elab_Tactic_Omega_Problem_addInequality__proof___closed__3; -x_35 = l_Lean_mkApp4(x_34, x_33, x_22, x_13, x_17); -lean_ctor_set(x_15, 0, x_35); -return x_15; +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_35 = l_Int_toNat(x_1); +x_36 = l_Lean_instToExprInt_mkNat(x_35); +x_37 = l_Lean_Elab_Tactic_Omega_Problem_addInequality__proof___closed__3; +x_38 = l_Lean_mkApp4(x_37, x_36, x_25, x_15, x_20); +lean_ctor_set(x_18, 0, x_38); +return x_18; } } else { -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; -x_36 = lean_ctor_get(x_15, 0); -x_37 = lean_ctor_get(x_15, 1); -lean_inc(x_37); -lean_inc(x_36); -lean_dec(x_15); -x_38 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1___closed__1; -x_39 = lean_int_dec_le(x_38, x_1); -x_40 = l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___lambda__1___closed__10; -x_41 = l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___lambda__1___closed__14; -x_42 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1(x_40, x_41, x_2); -if (x_39 == 0) -{ -lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_43 = lean_int_neg(x_1); -x_44 = l_Int_toNat(x_43); -lean_dec(x_43); -x_45 = l_Lean_instToExprInt_mkNat(x_44); -x_46 = l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___lambda__1___closed__16; -x_47 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1___closed__10; -x_48 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1___closed__13; -x_49 = l_Lean_mkApp3(x_46, x_47, x_48, x_45); -x_50 = l_Lean_Elab_Tactic_Omega_Problem_addInequality__proof___closed__3; -x_51 = l_Lean_mkApp4(x_50, x_49, x_42, x_13, x_36); -x_52 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_52, 0, x_51); -lean_ctor_set(x_52, 1, x_37); -return x_52; +lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_39 = lean_ctor_get(x_18, 0); +x_40 = lean_ctor_get(x_18, 1); +lean_inc(x_40); +lean_inc(x_39); +lean_dec(x_18); +x_41 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1___closed__1; +x_42 = lean_int_dec_le(x_41, x_1); +x_43 = l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___lambda__1___closed__10; +x_44 = l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___lambda__1___closed__14; +x_45 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1(x_43, x_44, x_2); +if (x_42 == 0) +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_46 = lean_int_neg(x_1); +x_47 = l_Int_toNat(x_46); +lean_dec(x_46); +x_48 = l_Lean_instToExprInt_mkNat(x_47); +x_49 = l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___lambda__1___closed__16; +x_50 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1___closed__10; +x_51 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1___closed__13; +x_52 = l_Lean_mkApp3(x_49, x_50, x_51, x_48); +x_53 = l_Lean_Elab_Tactic_Omega_Problem_addInequality__proof___closed__3; +x_54 = l_Lean_mkApp4(x_53, x_52, x_45, x_15, x_39); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_40); +return x_55; } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_53 = l_Int_toNat(x_1); -x_54 = l_Lean_instToExprInt_mkNat(x_53); -x_55 = l_Lean_Elab_Tactic_Omega_Problem_addInequality__proof___closed__3; -x_56 = l_Lean_mkApp4(x_55, x_54, x_42, x_13, x_36); -x_57 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_57, 0, x_56); -lean_ctor_set(x_57, 1, x_37); -return x_57; +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_56 = l_Int_toNat(x_1); +x_57 = l_Lean_instToExprInt_mkNat(x_56); +x_58 = l_Lean_Elab_Tactic_Omega_Problem_addInequality__proof___closed__3; +x_59 = l_Lean_mkApp4(x_58, x_57, x_45, x_15, x_39); +x_60 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_60, 0, x_59); +lean_ctor_set(x_60, 1, x_40); +return x_60; } } } else { -uint8_t x_58; -lean_dec(x_13); -x_58 = !lean_is_exclusive(x_15); -if (x_58 == 0) +uint8_t x_61; +lean_dec(x_15); +x_61 = !lean_is_exclusive(x_18); +if (x_61 == 0) { -return x_15; +return x_18; } else { -lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_59 = lean_ctor_get(x_15, 0); -x_60 = lean_ctor_get(x_15, 1); -lean_inc(x_60); -lean_inc(x_59); -lean_dec(x_15); -x_61 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_61, 0, x_59); -lean_ctor_set(x_61, 1, x_60); -return x_61; +lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_62 = lean_ctor_get(x_18, 0); +x_63 = lean_ctor_get(x_18, 1); +lean_inc(x_63); +lean_inc(x_62); +lean_dec(x_18); +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_62; +uint8_t x_65; +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_62 = !lean_is_exclusive(x_12); -if (x_62 == 0) +x_65 = !lean_is_exclusive(x_14); +if (x_65 == 0) { -return x_12; +return x_14; } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_63 = lean_ctor_get(x_12, 0); -x_64 = lean_ctor_get(x_12, 1); -lean_inc(x_64); -lean_inc(x_63); -lean_dec(x_12); -x_65 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_65, 0, x_63); -lean_ctor_set(x_65, 1, x_64); -return x_65; +lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_66 = lean_ctor_get(x_14, 0); +x_67 = lean_ctor_get(x_14, 1); +lean_inc(x_67); +lean_inc(x_66); +lean_dec(x_14); +x_68 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_68, 0, x_66); +lean_ctor_set(x_68, 1, x_67); +return x_68; } } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_addInequality__proof___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_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_Omega_Problem_addInequality__proof___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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_12; -x_12 = l_Lean_Elab_Tactic_Omega_Problem_addInequality__proof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +uint8_t x_14; lean_object* x_15; +x_14 = lean_unbox(x_7); +lean_dec(x_7); +x_15 = l_Lean_Elab_Tactic_Omega_Problem_addInequality__proof(x_1, x_2, x_3, x_4, x_5, x_6, x_14, x_8, x_9, x_10, x_11, x_12, x_13); lean_dec(x_2); lean_dec(x_1); -return x_12; +return x_15; } } static lean_object* _init_l_Lean_Elab_Tactic_Omega_Problem_addEquality__proof___closed__1() { @@ -10685,199 +10789,203 @@ x_3 = l_Lean_Expr_const___override(x_2, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_addEquality__proof(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_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_Omega_Problem_addEquality__proof(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_object* x_12, lean_object* x_13) { _start: { -lean_object* x_12; +lean_object* 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); -x_12 = l_Lean_Elab_Tactic_Omega_atomsCoeffs___rarg(x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_12) == 0) -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; -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_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); -if (lean_obj_tag(x_15) == 0) +x_14 = l_Lean_Elab_Tactic_Omega_atomsCoeffs___rarg(x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_14) == 0) { -uint8_t x_16; -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* 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 = lean_box(x_7); +x_18 = lean_apply_10(x_3, x_4, x_5, x_6, x_17, x_8, x_9, x_10, x_11, x_12, x_16); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_17; lean_object* x_18; uint8_t x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_17 = lean_ctor_get(x_15, 0); -x_18 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1___closed__1; -x_19 = lean_int_dec_le(x_18, x_1); -x_20 = l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___lambda__1___closed__10; -x_21 = l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___lambda__1___closed__14; -x_22 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1(x_20, x_21, x_2); +uint8_t x_19; +x_19 = !lean_is_exclusive(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; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_23 = lean_int_neg(x_1); -x_24 = l_Int_toNat(x_23); -lean_dec(x_23); -x_25 = l_Lean_instToExprInt_mkNat(x_24); -x_26 = l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___lambda__1___closed__16; -x_27 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1___closed__10; -x_28 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1___closed__13; -x_29 = l_Lean_mkApp3(x_26, x_27, x_28, x_25); -x_30 = l_Lean_Elab_Tactic_Omega_Problem_addEquality__proof___closed__3; -x_31 = l_Lean_mkApp4(x_30, x_29, x_22, x_13, x_17); -lean_ctor_set(x_15, 0, x_31); -return x_15; +lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_20 = lean_ctor_get(x_18, 0); +x_21 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1___closed__1; +x_22 = lean_int_dec_le(x_21, x_1); +x_23 = l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___lambda__1___closed__10; +x_24 = l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___lambda__1___closed__14; +x_25 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1(x_23, x_24, x_2); +if (x_22 == 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 = lean_int_neg(x_1); +x_27 = l_Int_toNat(x_26); +lean_dec(x_26); +x_28 = l_Lean_instToExprInt_mkNat(x_27); +x_29 = l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___lambda__1___closed__16; +x_30 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1___closed__10; +x_31 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1___closed__13; +x_32 = l_Lean_mkApp3(x_29, x_30, x_31, x_28); +x_33 = l_Lean_Elab_Tactic_Omega_Problem_addEquality__proof___closed__3; +x_34 = l_Lean_mkApp4(x_33, x_32, x_25, x_15, x_20); +lean_ctor_set(x_18, 0, x_34); +return x_18; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_32 = l_Int_toNat(x_1); -x_33 = l_Lean_instToExprInt_mkNat(x_32); -x_34 = l_Lean_Elab_Tactic_Omega_Problem_addEquality__proof___closed__3; -x_35 = l_Lean_mkApp4(x_34, x_33, x_22, x_13, x_17); -lean_ctor_set(x_15, 0, x_35); -return x_15; +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_35 = l_Int_toNat(x_1); +x_36 = l_Lean_instToExprInt_mkNat(x_35); +x_37 = l_Lean_Elab_Tactic_Omega_Problem_addEquality__proof___closed__3; +x_38 = l_Lean_mkApp4(x_37, x_36, x_25, x_15, x_20); +lean_ctor_set(x_18, 0, x_38); +return x_18; } } else { -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; -x_36 = lean_ctor_get(x_15, 0); -x_37 = lean_ctor_get(x_15, 1); -lean_inc(x_37); -lean_inc(x_36); -lean_dec(x_15); -x_38 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1___closed__1; -x_39 = lean_int_dec_le(x_38, x_1); -x_40 = l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___lambda__1___closed__10; -x_41 = l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___lambda__1___closed__14; -x_42 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1(x_40, x_41, x_2); -if (x_39 == 0) -{ -lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_43 = lean_int_neg(x_1); -x_44 = l_Int_toNat(x_43); -lean_dec(x_43); -x_45 = l_Lean_instToExprInt_mkNat(x_44); -x_46 = l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___lambda__1___closed__16; -x_47 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1___closed__10; -x_48 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1___closed__13; -x_49 = l_Lean_mkApp3(x_46, x_47, x_48, x_45); -x_50 = l_Lean_Elab_Tactic_Omega_Problem_addEquality__proof___closed__3; -x_51 = l_Lean_mkApp4(x_50, x_49, x_42, x_13, x_36); -x_52 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_52, 0, x_51); -lean_ctor_set(x_52, 1, x_37); -return x_52; +lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_39 = lean_ctor_get(x_18, 0); +x_40 = lean_ctor_get(x_18, 1); +lean_inc(x_40); +lean_inc(x_39); +lean_dec(x_18); +x_41 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1___closed__1; +x_42 = lean_int_dec_le(x_41, x_1); +x_43 = l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___lambda__1___closed__10; +x_44 = l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___lambda__1___closed__14; +x_45 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1(x_43, x_44, x_2); +if (x_42 == 0) +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_46 = lean_int_neg(x_1); +x_47 = l_Int_toNat(x_46); +lean_dec(x_46); +x_48 = l_Lean_instToExprInt_mkNat(x_47); +x_49 = l_Lean_Elab_Tactic_Omega_instToExprLinearCombo___lambda__1___closed__16; +x_50 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1___closed__10; +x_51 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1___closed__13; +x_52 = l_Lean_mkApp3(x_49, x_50, x_51, x_48); +x_53 = l_Lean_Elab_Tactic_Omega_Problem_addEquality__proof___closed__3; +x_54 = l_Lean_mkApp4(x_53, x_52, x_45, x_15, x_39); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_40); +return x_55; } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_53 = l_Int_toNat(x_1); -x_54 = l_Lean_instToExprInt_mkNat(x_53); -x_55 = l_Lean_Elab_Tactic_Omega_Problem_addEquality__proof___closed__3; -x_56 = l_Lean_mkApp4(x_55, x_54, x_42, x_13, x_36); -x_57 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_57, 0, x_56); -lean_ctor_set(x_57, 1, x_37); -return x_57; +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_56 = l_Int_toNat(x_1); +x_57 = l_Lean_instToExprInt_mkNat(x_56); +x_58 = l_Lean_Elab_Tactic_Omega_Problem_addEquality__proof___closed__3; +x_59 = l_Lean_mkApp4(x_58, x_57, x_45, x_15, x_39); +x_60 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_60, 0, x_59); +lean_ctor_set(x_60, 1, x_40); +return x_60; } } } else { -uint8_t x_58; -lean_dec(x_13); -x_58 = !lean_is_exclusive(x_15); -if (x_58 == 0) +uint8_t x_61; +lean_dec(x_15); +x_61 = !lean_is_exclusive(x_18); +if (x_61 == 0) { -return x_15; +return x_18; } else { -lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_59 = lean_ctor_get(x_15, 0); -x_60 = lean_ctor_get(x_15, 1); -lean_inc(x_60); -lean_inc(x_59); -lean_dec(x_15); -x_61 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_61, 0, x_59); -lean_ctor_set(x_61, 1, x_60); -return x_61; +lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_62 = lean_ctor_get(x_18, 0); +x_63 = lean_ctor_get(x_18, 1); +lean_inc(x_63); +lean_inc(x_62); +lean_dec(x_18); +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_62; +uint8_t x_65; +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_62 = !lean_is_exclusive(x_12); -if (x_62 == 0) +x_65 = !lean_is_exclusive(x_14); +if (x_65 == 0) { -return x_12; +return x_14; } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_63 = lean_ctor_get(x_12, 0); -x_64 = lean_ctor_get(x_12, 1); -lean_inc(x_64); -lean_inc(x_63); -lean_dec(x_12); -x_65 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_65, 0, x_63); -lean_ctor_set(x_65, 1, x_64); -return x_65; +lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_66 = lean_ctor_get(x_14, 0); +x_67 = lean_ctor_get(x_14, 1); +lean_inc(x_67); +lean_inc(x_66); +lean_dec(x_14); +x_68 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_68, 0, x_66); +lean_ctor_set(x_68, 1, x_67); +return x_68; } } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_addEquality__proof___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_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_Omega_Problem_addEquality__proof___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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_12; -x_12 = l_Lean_Elab_Tactic_Omega_Problem_addEquality__proof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +uint8_t x_14; lean_object* x_15; +x_14 = lean_unbox(x_7); +lean_dec(x_7); +x_15 = l_Lean_Elab_Tactic_Omega_Problem_addEquality__proof(x_1, x_2, x_3, x_4, x_5, x_6, x_14, x_8, x_9, x_10, x_11, x_12, x_13); lean_dec(x_2); lean_dec(x_1); -return x_12; +return x_15; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_addInequality___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_Tactic_Omega_Problem_addInequality___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) { _start: { -uint8_t x_8; lean_object* x_9; lean_object* x_10; -x_8 = 0; -x_9 = lean_box(0); -x_10 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkFreshExprMVarImpl(x_1, x_8, x_9, x_3, x_4, x_5, x_6, x_7); -return x_10; +uint8_t x_10; lean_object* x_11; lean_object* x_12; +x_10 = 0; +x_11 = lean_box(0); +x_12 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkFreshExprMVarImpl(x_1, x_10, x_11, x_5, x_6, x_7, x_8, x_9); +return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_addInequality___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_addInequality___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, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -uint8_t x_10; lean_object* x_11; -x_10 = 0; -x_11 = l_Lean_Meta_mkSorry(x_1, x_10, x_5, x_6, x_7, x_8, x_9); -return x_11; +uint8_t x_12; lean_object* x_13; +x_12 = 0; +x_13 = l_Lean_Meta_mkSorry(x_1, x_12, x_7, x_8, x_9, x_10, x_11); +return x_13; } } static lean_object* _init_l_Lean_Elab_Tactic_Omega_Problem_addInequality___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_Problem_addInequality___lambda__2___boxed), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_Problem_addInequality___lambda__2___boxed), 11, 0); return x_1; } } @@ -10943,17 +11051,17 @@ x_29 = l_Lean_Elab_Tactic_Omega_Fact_tidy(x_20); if (lean_obj_tag(x_4) == 0) { lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_30 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_Problem_addInequality___lambda__1___boxed), 7, 1); +x_30 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_Problem_addInequality___lambda__1___boxed), 9, 1); lean_closure_set(x_30, 0, x_16); x_31 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); lean_closure_set(x_31, 0, x_30); x_32 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); lean_closure_set(x_32, 0, x_31); x_33 = l_Lean_Elab_Tactic_Omega_Problem_addInequality___closed__1; -x_34 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg), 10, 2); +x_34 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg___boxed), 12, 2); lean_closure_set(x_34, 0, x_32); lean_closure_set(x_34, 1, x_33); -x_35 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_Problem_addInequality__proof___boxed), 11, 3); +x_35 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_Problem_addInequality__proof___boxed), 13, 3); lean_closure_set(x_35, 0, x_2); lean_closure_set(x_35, 1, x_3); lean_closure_set(x_35, 2, x_34); @@ -10968,7 +11076,7 @@ lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; x_38 = lean_ctor_get(x_4, 0); lean_inc(x_38); lean_dec(x_4); -x_39 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_Problem_addInequality__proof___boxed), 11, 3); +x_39 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_Problem_addInequality__proof___boxed), 13, 3); lean_closure_set(x_39, 0, x_2); lean_closure_set(x_39, 1, x_3); lean_closure_set(x_39, 2, x_38); @@ -10986,17 +11094,17 @@ x_42 = l_Lean_Elab_Tactic_Omega_Fact_tidy(x_20); if (lean_obj_tag(x_4) == 0) { lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_43 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_Problem_addInequality___lambda__1___boxed), 7, 1); +x_43 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_Problem_addInequality___lambda__1___boxed), 9, 1); lean_closure_set(x_43, 0, x_16); x_44 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); lean_closure_set(x_44, 0, x_43); x_45 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); lean_closure_set(x_45, 0, x_44); x_46 = l_Lean_Elab_Tactic_Omega_Problem_addInequality___closed__1; -x_47 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg), 10, 2); +x_47 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg___boxed), 12, 2); lean_closure_set(x_47, 0, x_45); lean_closure_set(x_47, 1, x_46); -x_48 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_Problem_addInequality__proof___boxed), 11, 3); +x_48 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_Problem_addInequality__proof___boxed), 13, 3); lean_closure_set(x_48, 0, x_2); lean_closure_set(x_48, 1, x_3); lean_closure_set(x_48, 2, x_47); @@ -11019,7 +11127,7 @@ lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean x_52 = lean_ctor_get(x_4, 0); lean_inc(x_52); lean_dec(x_4); -x_53 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_Problem_addInequality__proof___boxed), 11, 3); +x_53 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_Problem_addInequality__proof___boxed), 13, 3); lean_closure_set(x_53, 0, x_2); lean_closure_set(x_53, 1, x_3); lean_closure_set(x_53, 2, x_52); @@ -11039,27 +11147,33 @@ return x_56; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_addInequality___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_Elab_Tactic_Omega_Problem_addInequality___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_8; -x_8 = l_Lean_Elab_Tactic_Omega_Problem_addInequality___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +uint8_t x_10; lean_object* x_11; +x_10 = lean_unbox(x_3); +lean_dec(x_3); +x_11 = l_Lean_Elab_Tactic_Omega_Problem_addInequality___lambda__1(x_1, x_2, x_10, 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_2); -return x_8; +return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_addInequality___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_Elab_Tactic_Omega_Problem_addInequality___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_10; -x_10 = l_Lean_Elab_Tactic_Omega_Problem_addInequality___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +uint8_t x_12; lean_object* x_13; +x_12 = lean_unbox(x_5); +lean_dec(x_5); +x_13 = l_Lean_Elab_Tactic_Omega_Problem_addInequality___lambda__2(x_1, x_2, x_3, x_4, x_12, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_10; +return x_13; } } static lean_object* _init_l_Lean_Elab_Tactic_Omega_Problem_addEquality___closed__1() { @@ -11067,7 +11181,7 @@ static lean_object* _init_l_Lean_Elab_Tactic_Omega_Problem_addEquality___closed_ { lean_object* x_1; lean_object* x_2; x_1 = lean_box(0); -x_2 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_Problem_addInequality___lambda__1___boxed), 7, 1); +x_2 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_Problem_addInequality___lambda__1___boxed), 9, 1); lean_closure_set(x_2, 0, x_1); return x_2; } @@ -11098,7 +11212,7 @@ static lean_object* _init_l_Lean_Elab_Tactic_Omega_Problem_addEquality___closed_ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Tactic_Omega_Problem_addEquality___closed__3; x_2 = l_Lean_Elab_Tactic_Omega_Problem_addInequality___closed__1; -x_3 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg), 10, 2); +x_3 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg___boxed), 12, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); return x_3; @@ -11167,7 +11281,7 @@ if (lean_obj_tag(x_4) == 0) { lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; x_29 = l_Lean_Elab_Tactic_Omega_Problem_addEquality___closed__4; -x_30 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_Problem_addEquality__proof___boxed), 11, 3); +x_30 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_Problem_addEquality__proof___boxed), 13, 3); lean_closure_set(x_30, 0, x_2); lean_closure_set(x_30, 1, x_3); lean_closure_set(x_30, 2, x_29); @@ -11182,7 +11296,7 @@ lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; x_33 = lean_ctor_get(x_4, 0); lean_inc(x_33); lean_dec(x_4); -x_34 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_Problem_addEquality__proof___boxed), 11, 3); +x_34 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_Problem_addEquality__proof___boxed), 13, 3); lean_closure_set(x_34, 0, x_2); lean_closure_set(x_34, 1, x_3); lean_closure_set(x_34, 2, x_33); @@ -11201,7 +11315,7 @@ if (lean_obj_tag(x_4) == 0) { lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; x_38 = l_Lean_Elab_Tactic_Omega_Problem_addEquality___closed__4; -x_39 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_Problem_addEquality__proof___boxed), 11, 3); +x_39 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_Problem_addEquality__proof___boxed), 13, 3); lean_closure_set(x_39, 0, x_2); lean_closure_set(x_39, 1, x_3); lean_closure_set(x_39, 2, x_38); @@ -11224,7 +11338,7 @@ lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean x_43 = lean_ctor_get(x_4, 0); lean_inc(x_43); lean_dec(x_4); -x_44 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_Problem_addEquality__proof___boxed), 11, 3); +x_44 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_Problem_addEquality__proof___boxed), 13, 3); lean_closure_set(x_44, 0, x_2); lean_closure_set(x_44, 1, x_3); lean_closure_set(x_44, 2, x_43); @@ -15625,76 +15739,79 @@ lean_dec(x_4); return x_9; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_runOmega___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_Elab_Tactic_Omega_Problem_runOmega___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_object* x_11, lean_object* x_12) { _start: { -uint8_t x_11; +uint8_t x_13; lean_dec(x_2); -x_11 = lean_ctor_get_uint8(x_1, sizeof(void*)*7); -if (x_11 == 0) +x_13 = lean_ctor_get_uint8(x_1, sizeof(void*)*7); +if (x_13 == 0) { -lean_object* x_12; +lean_object* x_14; +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_1); -lean_ctor_set(x_12, 1, x_10); -return x_12; +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_1); +lean_ctor_set(x_14, 1, x_12); +return x_14; } else { -lean_object* x_13; +lean_object* x_15; +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_13 = l_Lean_Elab_Tactic_Omega_Problem_solveEqualities(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_13) == 0) +x_15 = l_Lean_Elab_Tactic_Omega_Problem_solveEqualities(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); -x_16 = l_Lean_Elab_Tactic_Omega_Problem_elimination(x_14, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_15); -return x_16; +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = l_Lean_Elab_Tactic_Omega_Problem_elimination(x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_17); +return x_18; } else { -uint8_t x_17; +uint8_t x_19; +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_17 = !lean_is_exclusive(x_13); -if (x_17 == 0) +x_19 = !lean_is_exclusive(x_15); +if (x_19 == 0) { -return x_13; +return x_15; } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_13, 0); -x_19 = lean_ctor_get(x_13, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_13); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_18); -lean_ctor_set(x_20, 1, x_19); -return x_20; +lean_object* x_20; lean_object* x_21; lean_object* x_22; +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_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +return x_22; } } } @@ -15805,118 +15922,118 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_runOmega(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_runOmega(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) { _start: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_10 = l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Core___hyg_5____closed__2; -x_11 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_10, x_2, x_3, x_4, 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); -lean_dec(x_12); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_11, 1); +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_12 = l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Core___hyg_5____closed__2; +x_13 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_12, x_2, x_3, x_4, 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); -lean_dec(x_11); -x_15 = lean_box(0); -x_16 = l_Lean_Elab_Tactic_Omega_Problem_runOmega___lambda__1(x_1, x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_14); -return x_16; +x_15 = lean_unbox(x_14); +lean_dec(x_14); +if (x_15 == 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_dec(x_13); +x_17 = lean_box(0); +x_18 = l_Lean_Elab_Tactic_Omega_Problem_runOmega___lambda__1(x_1, x_17, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_16); +return x_18; } else { -uint8_t x_17; -x_17 = lean_ctor_get_uint8(x_1, sizeof(void*)*7); -if (x_17 == 0) +uint8_t x_19; +x_19 = lean_ctor_get_uint8(x_1, sizeof(void*)*7); +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; -x_18 = lean_ctor_get(x_11, 1); -lean_inc(x_18); -lean_dec(x_11); -x_19 = l_Lean_Elab_Tactic_Omega_Problem_runOmega___closed__6; -x_20 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_10, x_19, x_2, x_3, x_4, x_5, x_6, x_7, x_8, 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_Tactic_Omega_Problem_runOmega___lambda__1(x_1, x_21, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_22); -return x_23; +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_20 = lean_ctor_get(x_13, 1); +lean_inc(x_20); +lean_dec(x_13); +x_21 = l_Lean_Elab_Tactic_Omega_Problem_runOmega___closed__6; +x_22 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_12, x_21, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_20); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +lean_dec(x_22); +x_25 = l_Lean_Elab_Tactic_Omega_Problem_runOmega___lambda__1(x_1, x_23, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_24); +return x_25; } else { -lean_object* x_24; uint8_t x_25; -x_24 = lean_ctor_get(x_11, 1); -lean_inc(x_24); -lean_dec(x_11); -x_25 = l_Lean_Elab_Tactic_Omega_Problem_isEmpty(x_1); -if (x_25 == 0) -{ -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; 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_26 = lean_ctor_get(x_1, 2); +lean_object* x_26; uint8_t x_27; +x_26 = lean_ctor_get(x_13, 1); lean_inc(x_26); -x_27 = l_Lean_HashMap_toList___at_Lean_Elab_Tactic_Omega_Problem_instToStringProblem___spec__1(x_26); -x_28 = lean_box(0); -x_29 = l_List_mapTR_loop___at_Lean_Elab_Tactic_Omega_Problem_instToStringProblem___spec__4(x_27, x_28); -x_30 = l___private_Lean_Elab_Tactic_Omega_Core_0__Lean_Elab_Tactic_Omega_Justification_bullet___closed__1; -x_31 = l_String_intercalate(x_30, x_29); -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_Elab_Tactic_Omega_Problem_runOmega___closed__2; -x_35 = lean_alloc_ctor(7, 2, 0); +lean_dec(x_13); +x_27 = l_Lean_Elab_Tactic_Omega_Problem_isEmpty(x_1); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_28 = lean_ctor_get(x_1, 2); +lean_inc(x_28); +x_29 = l_Lean_HashMap_toList___at_Lean_Elab_Tactic_Omega_Problem_instToStringProblem___spec__1(x_28); +x_30 = lean_box(0); +x_31 = l_List_mapTR_loop___at_Lean_Elab_Tactic_Omega_Problem_instToStringProblem___spec__4(x_29, x_30); +x_32 = l___private_Lean_Elab_Tactic_Omega_Core_0__Lean_Elab_Tactic_Omega_Justification_bullet___closed__1; +x_33 = l_String_intercalate(x_32, x_31); +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_Std_Range_forIn_loop___at_Lean_Elab_Tactic_Omega_Problem_fourierMotzkinSelect___spec__1___closed__3; +x_36 = l_Lean_Elab_Tactic_Omega_Problem_runOmega___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 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_10, x_37, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_24); -x_39 = lean_ctor_get(x_38, 0); -lean_inc(x_39); -x_40 = lean_ctor_get(x_38, 1); -lean_inc(x_40); -lean_dec(x_38); -x_41 = l_Lean_Elab_Tactic_Omega_Problem_runOmega___lambda__1(x_1, x_39, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_40); -return x_41; +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_35); +x_38 = l_Std_Range_forIn_loop___at_Lean_Elab_Tactic_Omega_Problem_fourierMotzkinSelect___spec__1___closed__3; +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_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_12, x_39, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_26); +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_Tactic_Omega_Problem_runOmega___lambda__1(x_1, x_41, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_42); +return x_43; } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_42 = l_Lean_Elab_Tactic_Omega_Problem_runOmega___closed__10; -x_43 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_10, x_42, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_24); -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_Elab_Tactic_Omega_Problem_runOmega___lambda__1(x_1, x_44, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_45); -return x_46; +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_44 = l_Lean_Elab_Tactic_Omega_Problem_runOmega___closed__10; +x_45 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_12, x_44, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_26); +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 = l_Lean_Elab_Tactic_Omega_Problem_runOmega___lambda__1(x_1, x_46, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_47); +return x_48; } } } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_elimination___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_Elab_Tactic_Omega_Problem_elimination___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_object* x_11, lean_object* x_12) { _start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_dec(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); -x_11 = l_Lean_Elab_Tactic_Omega_Problem_fourierMotzkin(x_1, 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_Elab_Tactic_Omega_Problem_runOmega(x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); -return x_14; +x_13 = l_Lean_Elab_Tactic_Omega_Problem_fourierMotzkin(x_1, x_8, x_9, x_10, x_11, x_12); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_Elab_Tactic_Omega_Problem_runOmega(x_14, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_15); +return x_16; } } static lean_object* _init_l_Lean_Elab_Tactic_Omega_Problem_elimination___closed__1() { @@ -15936,100 +16053,142 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_elimination(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_elimination(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) { _start: { -uint8_t x_10; -x_10 = lean_ctor_get_uint8(x_1, sizeof(void*)*7); -if (x_10 == 0) +uint8_t x_12; +x_12 = lean_ctor_get_uint8(x_1, sizeof(void*)*7); +if (x_12 == 0) { -lean_object* x_11; +lean_object* x_13; +lean_dec(x_10); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_11 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_11, 0, x_1); -lean_ctor_set(x_11, 1, x_9); -return x_11; +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_1); +lean_ctor_set(x_13, 1, x_11); +return x_13; } else { -uint8_t x_12; -x_12 = l_Lean_Elab_Tactic_Omega_Problem_isEmpty(x_1); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_13 = l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Core___hyg_5____closed__2; -x_14 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_13, x_2, x_3, x_4, 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_unbox(x_15); -lean_dec(x_15); -if (x_16 == 0) +uint8_t x_14; +x_14 = l_Lean_Elab_Tactic_Omega_Problem_isEmpty(x_1); +if (x_14 == 0) { -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_ctor_get(x_14, 1); +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_15 = l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Core___hyg_5____closed__2; +x_16 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_17 = lean_ctor_get(x_16, 0); lean_inc(x_17); -lean_dec(x_14); -x_18 = lean_box(0); -x_19 = l_Lean_Elab_Tactic_Omega_Problem_elimination___lambda__1(x_1, x_18, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_17); -return x_19; +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; +x_19 = lean_ctor_get(x_16, 1); +lean_inc(x_19); +lean_dec(x_16); +x_20 = lean_box(0); +x_21 = l_Lean_Elab_Tactic_Omega_Problem_elimination___lambda__1(x_1, x_20, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_19); +return x_21; } 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_14, 1); -lean_inc(x_20); -lean_dec(x_14); -x_21 = lean_ctor_get(x_1, 2); -lean_inc(x_21); -x_22 = l_Lean_HashMap_toList___at_Lean_Elab_Tactic_Omega_Problem_instToStringProblem___spec__1(x_21); -x_23 = lean_box(0); -x_24 = l_List_mapTR_loop___at_Lean_Elab_Tactic_Omega_Problem_instToStringProblem___spec__4(x_22, x_23); -x_25 = l___private_Lean_Elab_Tactic_Omega_Core_0__Lean_Elab_Tactic_Omega_Justification_bullet___closed__1; -x_26 = l_String_intercalate(x_25, x_24); -x_27 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_27, 0, x_26); -x_28 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_28, 0, x_27); -x_29 = l_Lean_Elab_Tactic_Omega_Problem_elimination___closed__2; -x_30 = lean_alloc_ctor(7, 2, 0); +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_22 = lean_ctor_get(x_16, 1); +lean_inc(x_22); +lean_dec(x_16); +x_23 = lean_ctor_get(x_1, 2); +lean_inc(x_23); +x_24 = l_Lean_HashMap_toList___at_Lean_Elab_Tactic_Omega_Problem_instToStringProblem___spec__1(x_23); +x_25 = lean_box(0); +x_26 = l_List_mapTR_loop___at_Lean_Elab_Tactic_Omega_Problem_instToStringProblem___spec__4(x_24, x_25); +x_27 = l___private_Lean_Elab_Tactic_Omega_Core_0__Lean_Elab_Tactic_Omega_Justification_bullet___closed__1; +x_28 = l_String_intercalate(x_27, x_26); +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); -lean_ctor_set(x_30, 1, x_28); -x_31 = l_Std_Range_forIn_loop___at_Lean_Elab_Tactic_Omega_Problem_fourierMotzkinSelect___spec__1___closed__3; +x_31 = l_Lean_Elab_Tactic_Omega_Problem_elimination___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 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_13, x_32, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_20); -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_Lean_Elab_Tactic_Omega_Problem_elimination___lambda__1(x_1, x_34, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_35); -return x_36; +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_30); +x_33 = l_Std_Range_forIn_loop___at_Lean_Elab_Tactic_Omega_Problem_fourierMotzkinSelect___spec__1___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 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_15, x_34, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_22); +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_Elab_Tactic_Omega_Problem_elimination___lambda__1(x_1, x_36, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_37); +return x_38; } } else { -lean_object* x_37; +lean_object* x_39; +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_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_1); -lean_ctor_set(x_37, 1, x_9); -return x_37; +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_1); +lean_ctor_set(x_39, 1, x_11); +return x_39; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_runOmega___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; lean_object* x_14; +x_13 = lean_unbox(x_6); +lean_dec(x_6); +x_14 = l_Lean_Elab_Tactic_Omega_Problem_runOmega___lambda__1(x_1, x_2, x_3, x_4, x_5, x_13, x_7, x_8, x_9, x_10, x_11, x_12); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_runOmega___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_5); +lean_dec(x_5); +x_13 = l_Lean_Elab_Tactic_Omega_Problem_runOmega(x_1, x_2, x_3, x_4, x_12, x_6, x_7, x_8, x_9, x_10, x_11); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_elimination___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; lean_object* x_14; +x_13 = lean_unbox(x_6); +lean_dec(x_6); +x_14 = l_Lean_Elab_Tactic_Omega_Problem_elimination___lambda__1(x_1, x_2, x_3, x_4, x_5, x_13, x_7, x_8, x_9, x_10, x_11, x_12); +return x_14; } } +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_Problem_elimination___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_5); +lean_dec(x_5); +x_13 = l_Lean_Elab_Tactic_Omega_Problem_elimination(x_1, x_2, x_3, x_4, x_12, x_6, x_7, x_8, x_9, x_10, x_11); +return x_13; } } lean_object* initialize_Init_Omega_Constraint(uint8_t builtin, lean_object*); @@ -16336,6 +16495,10 @@ l_Lean_Elab_Tactic_Omega_Justification_proof___rarg___closed__3 = _init_l_Lean_E lean_mark_persistent(l_Lean_Elab_Tactic_Omega_Justification_proof___rarg___closed__3); l_Lean_Elab_Tactic_Omega_Justification_proof___rarg___closed__4 = _init_l_Lean_Elab_Tactic_Omega_Justification_proof___rarg___closed__4(); lean_mark_persistent(l_Lean_Elab_Tactic_Omega_Justification_proof___rarg___closed__4); +l_Lean_Elab_Tactic_Omega_Justification_proof___rarg___closed__5 = _init_l_Lean_Elab_Tactic_Omega_Justification_proof___rarg___closed__5(); +lean_mark_persistent(l_Lean_Elab_Tactic_Omega_Justification_proof___rarg___closed__5); +l_Lean_Elab_Tactic_Omega_Justification_proof___rarg___closed__6 = _init_l_Lean_Elab_Tactic_Omega_Justification_proof___rarg___closed__6(); +lean_mark_persistent(l_Lean_Elab_Tactic_Omega_Justification_proof___rarg___closed__6); l___auto____x40_Lean_Elab_Tactic_Omega_Core___hyg_1659____closed__1 = _init_l___auto____x40_Lean_Elab_Tactic_Omega_Core___hyg_1659____closed__1(); lean_mark_persistent(l___auto____x40_Lean_Elab_Tactic_Omega_Core___hyg_1659____closed__1); l___auto____x40_Lean_Elab_Tactic_Omega_Core___hyg_1659____closed__2 = _init_l___auto____x40_Lean_Elab_Tactic_Omega_Core___hyg_1659____closed__2(); diff --git a/stage0/stdlib/Lean/Elab/Tactic/Omega/Frontend.c b/stage0/stdlib/Lean/Elab/Tactic/Omega/Frontend.c index 5db587e635df..d3a3cade5647 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Omega/Frontend.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Omega/Frontend.c @@ -23,22 +23,21 @@ static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___clo static lean_object* l_Lean_Elab_Tactic_Omega_mkCoordinateEvalAtomsEq___closed__3; lean_object* l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___closed__1; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_mkEvalRflProof___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_Omega_mkEvalRflProof___boxed(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_Omega_asLinearComboImpl_handleFinVal___closed__3; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__50; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__6___closed__2; static lean_object* l_Lean_Elab_Tactic_Omega_evalOmega___closed__1; -lean_object* l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___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_Omega_MetaProblem_addIntEquality___closed__4; lean_object* l_Lean_mkNatLit(lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__35; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__5; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__3___closed__4; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___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_Tactic_Omega_omegaImpl___lambda__3(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_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__2(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_Elab_Tactic_Omega_omegaTactic___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_mkEvalRflProof(lean_object*, 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_Omega_mkEvalRflProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__66; LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__20; @@ -51,8 +50,10 @@ static lean_object* l_Lean_Elab_Tactic_Omega_mkCoordinateEvalAtomsEq___closed__7 LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__14(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_Omega_asLinearComboImpl___lambda__5___closed__6; lean_object* l_Lean_Elab_CommandContextInfo_save___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo___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_Omega_asLinearCombo___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_Omega_elabOmegaConfig___spec__6(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_problem___default; static lean_object* l_Lean_Elab_Tactic_Omega_succ_x3f___closed__5; @@ -60,22 +61,23 @@ static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___c static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___closed__1; static lean_object* l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__13; static lean_object* l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Tactic_Omega_elabOmegaConfig___spec__1___closed__1; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___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_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___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_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__4; static lean_object* l_Lean_Elab_Tactic_Omega_mkCoordinateEvalAtomsEq___closed__6; -lean_object* l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(lean_object*, 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_Omega_asLinearComboImpl(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_Tactic_Omega_lookup___spec__14(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_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__4___closed__3; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_Omega_elabOmegaConfig___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_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___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_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___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*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleFinVal___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_cases_u2082(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___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_Tactic_Omega_asLinearComboImpl___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__9___closed__1; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__6; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__13___closed__1; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__12___closed__2; static lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___closed__11; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleFinVal___closed__7; static lean_object* l_Lean_Elab_Tactic_Omega_cases_u2082___closed__3; LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality___spec__1(lean_object*, lean_object*, lean_object*); @@ -86,7 +88,7 @@ static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___c static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__33; static lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__7___closed__2; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__27; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleFinVal(lean_object*, lean_object*, lean_object*, 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_Omega_asLinearComboImpl_handleFinVal(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*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__1; static lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__6___closed__4; lean_object* l_Lean_Meta_isProp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -97,7 +99,7 @@ static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___c lean_object* l_Lean_Meta_mkEqSymm(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashSet_insert___at_Lean_Elab_Tactic_Omega_asLinearComboImpl___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__30; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality___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_Elab_Term_synthesizeSyntheticMVars(uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkAppB(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__15___closed__1; @@ -106,6 +108,7 @@ static lean_object* l_Lean_Elab_Tactic_Omega_succ_x3f___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_omegaTactic___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___closed__10; static lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_14702____closed__3; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___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_Elab_Tactic_Omega_omegaTactic___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_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__1; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__13; @@ -120,9 +123,10 @@ static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___clo uint8_t l_Lean_Expr_isApp(lean_object*); lean_object* l_Lean_Expr_lit___override(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*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality___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_Omega_MetaProblem_addIntInequality___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*); extern lean_object* l_Lean_Elab_Tactic_tacticElabAttribute; -lean_object* l_Lean_Elab_Tactic_Omega_withoutModifyingState___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_Tactic_Omega_MetaProblem_addFact___boxed(lean_object*, 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_Omega_withoutModifyingState___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* l_Lean_MessageData_ofList(lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_problem___default___closed__5; lean_object* lean_array_push(lean_object*, lean_object*); @@ -140,10 +144,10 @@ LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_Elab_Tactic_Omeg static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___closed__2; lean_object* lean_mk_array(lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at_Lean_Elab_Tactic_Omega_MetaProblem_addFact___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_instantiateMVars___at_Lean_Elab_Tactic_Omega_MetaProblem_addFact___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_Omega_asLinearComboImpl_handleNatCast___closed__9; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleFinVal___closed__12; -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_Omega_splitDisjunction___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_throwError___at_Lean_Elab_Tactic_Omega_splitDisjunction___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_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__6; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_Omega_evalOmega(lean_object*); @@ -162,6 +166,7 @@ lean_object* l_List_mapTR_loop___at_Lean_Omega_LinearCombo_instToStringLinearCom static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__3; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__9; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__15; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___boxed(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_Omega_mkEvalRflProof___closed__8; static lean_object* l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1___closed__5; 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*); @@ -173,14 +178,14 @@ LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Elab_Tactic_Omega_asLinearCombo LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_Omega_evalOmega_declRange___closed__2; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__6___closed__5; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___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*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__20; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___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_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___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*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__10___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__18(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_Omega_MetaProblem_addFact___lambda__1___closed__53; 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*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___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_Tactic_Omega_MetaProblem_addIntEquality___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_ReaderT_instMonadReaderT___rarg(lean_object*); lean_object* l_Lean_Expr_fvarId_x21(lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__12; @@ -217,13 +222,13 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_instInhabitedMetaP static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__7___closed__5; lean_object* l_Lean_mkApp4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__15; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___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_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___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*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_omegaTactic___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__37; static lean_object* l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__5; uint8_t lean_string_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__45; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___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_Omega_MetaProblem_addIntEquality___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*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__33; lean_object* l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality___closed__2; @@ -245,6 +250,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_evalOmega___lambda__1___boxed( static lean_object* l_Lean_Elab_Tactic_Omega_evalUnsafe____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_6____closed__2; lean_object* l_Lean_Level_ofNat(lean_object*); lean_object* l_Lean_Expr_appArg_x21(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_processFacts___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_Omega_asLinearComboImpl_handleFinVal___closed__4; static lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__3___closed__2; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__4___closed__3; @@ -255,22 +261,23 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Omega_a lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMap_insert___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__4(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__19___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality(lean_object*, lean_object*, 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_Omega_MetaProblem_addIntInequality(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*); uint8_t l_Lean_Expr_hasMVar(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSavedPartialInfoContext___at_Lean_Elab_Tactic_Omega_elabOmegaConfig___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_Elab_Tactic_Omega_omegaImpl___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_Omega_omegaImpl___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*, lean_object*, lean_object*); static lean_object* l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1___closed__2; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__7; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_Omega_elabOmegaConfig___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_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___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_Tactic_Omega_asLinearComboImpl_rewrite___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___closed__3; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___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*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__7___boxed(lean_object**); static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__4___closed__1; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl(lean_object*, 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_Omega_omegaImpl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__29; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__4(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_Omega_evalOmega_declRange___closed__4; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___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*); static lean_object* l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__25; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__15; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__42; @@ -285,7 +292,7 @@ static lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omeg LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__3(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__36; static lean_object* l_Lean_Elab_Tactic_Omega_elabOmegaConfig___closed__3; -lean_object* l_Lean_Elab_Tactic_Omega_Problem_elimination(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Elab_Tactic_Omega_Problem_elimination(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_elabOmegaConfig___closed__8; lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_evalUnsafe____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_6____closed__4; @@ -295,7 +302,7 @@ static lean_object* l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__16; static lean_object* l_Lean_Elab_Tactic_Omega_elabOmegaConfig___closed__11; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__1___boxed(lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__2; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_mkCoordinateEvalAtomsEq(lean_object*, 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_Omega_mkCoordinateEvalAtomsEq(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_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_elabOmegaConfig___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__28; static lean_object* l_Lean_Elab_Tactic_Omega_mkCoordinateEvalAtomsEq___closed__8; @@ -306,7 +313,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_succ_x3f(lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__46; lean_object* lean_st_ref_take(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_evalUnsafe____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_6____closed__3; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___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*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__7___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_withoutModifyingStateWithInfoAndMessages___at_Lean_Elab_Tactic_Omega_elabOmegaConfig___spec__13(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*); @@ -314,7 +321,7 @@ static lean_object* l_Lean_Elab_Tactic_Omega_evalUnsafe____x40_Lean_Elab_Tactic_ static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__14___closed__1; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__36; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__34; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___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_Elab_Tactic_Omega_asLinearComboImpl___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*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__16___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__4___closed__2; @@ -328,10 +335,10 @@ static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___c static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__1___closed__3; lean_object* l_List_mapTR_loop___at_Lean_MessageData_instCoeListExprMessageData___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__64; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast(lean_object*, lean_object*, 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_Omega_asLinearComboImpl_handleNatCast(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_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_Elab_Tactic_Omega_omegaTactic___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__10; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(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_Omega_mkAtomLinearCombo(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__24; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_Omega_omegaTactic___spec__6(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -360,10 +367,10 @@ lean_object* l_Lean_PersistentArray_append___rarg(lean_object*, lean_object*); lean_object* l_Lean_Elab_Tactic_Omega_groundNat_x3f(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_PersistentArray_mapMAux___at_Lean_Elab_Tactic_Omega_elabOmegaConfig___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_Elab_Tactic_Omega_omegaImpl___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_Omega_omegaImpl___lambda__1(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*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__18; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__4___closed__7; -lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(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_Tactic_Omega_lookup___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_List_elem___at_Lean_CollectLevelParams_visitExpr___spec__2(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__26; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -372,7 +379,7 @@ static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___c static lean_object* l_Lean_Elab_Tactic_Omega_mkCoordinateEvalAtomsEq___closed__10; lean_object* l_Lean_Meta_getLevel(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_appArg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality___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_Elab_Tactic_Omega_MetaProblem_addIntInequality___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__4___closed__2; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__9; 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*); @@ -388,17 +395,18 @@ lean_object* l_Lean_Meta_mkEqTrans(lean_object*, lean_object*, lean_object*, lea LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_Omega_cases_u2082___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__5; 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_Lean_Elab_Tactic_Omega_splitDisjunction___boxed(lean_object*, lean_object*, lean_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_Omega_omegaTactic___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_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__57; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__58; uint8_t l_List_isEmpty___rarg(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___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* lean_st_mk_ref(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__16; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__26; lean_object* l_Lean_Meta_mkDecideProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__34; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___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_Elab_Tactic_Omega_MetaProblem_addIntEquality___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_asLinearCombo___closed__1; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__11; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__1___closed__4; @@ -407,17 +415,17 @@ static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___c lean_object* lean_array_to_list(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__23; lean_object* l_Lean_Expr_appFnCleanup(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at_Lean_Elab_Tactic_Omega_omegaImpl___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_MVarId_assign___at_Lean_Elab_Tactic_Omega_omegaImpl___spec__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*); static lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___closed__6; 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_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__7___closed__2; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_mkCoordinateEvalAtomsEq___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_Omega_mkCoordinateEvalAtomsEq___boxed(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_Omega_MetaProblem_pushNot___closed__2; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__2; static lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__6___closed__1; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearCombo___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_Tactic_Omega_asLinearCombo___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_instInhabited___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__31; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__19___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -428,7 +436,7 @@ static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___c static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__40; static lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__3___closed__1; extern lean_object* l_Lean_instInhabitedExpr; -LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__3___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* l_Lean_Elab_InfoTree_substitute(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__2___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__1; @@ -442,8 +450,8 @@ static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__14___ static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__20; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleFinVal___closed__14; lean_object* l_List_lengthTRAux___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearCombo___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_Elab_Tactic_Omega_asLinearComboImpl___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_Tactic_Omega_asLinearComboImpl___lambda__8___closed__30; static lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___closed__12; static lean_object* l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__6; @@ -454,6 +462,7 @@ static lean_object* l_Lean_Elab_Tactic_Omega_omegaTactic___lambda__3___closed__1 static lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__6___closed__2; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); static lean_object* l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1___closed__7; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__6___boxed(lean_object**); static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__39; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___closed__1; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__41; @@ -464,12 +473,13 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_elabOmegaConfig(lean_object*, LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_omegaTactic___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkOptionalNode(lean_object*); lean_object* l_Lean_Elab_Tactic_Omega_Problem_addEquality(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___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_Omega_asLinearComboImpl_rewrite___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_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_Omega_elabOmegaConfig___spec__7(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___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_AssocList_contains___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__5___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__2(lean_object*, lean_object*, uint8_t, lean_object*, 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_Elab_Tactic_Omega_MetaProblem_addFact___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_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__2___closed__1; @@ -479,11 +489,13 @@ static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__9___c static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__25; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__5___closed__2; static lean_object* l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__11; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_processFacts(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_Omega_MetaProblem_processFacts(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_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_Omega_elabOmegaConfig___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*); static lean_object* l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__2; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___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_Omega_asLinearComboImpl___lambda__8___closed__14; -lean_object* l_Lean_Elab_Tactic_Omega_atomsCoeffs___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Elab_Tactic_Omega_atomsCoeffs___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__3___boxed(lean_object**); static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__50; static lean_object* l_Lean_Elab_Tactic_Omega_elabOmegaConfig___closed__9; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__8; @@ -491,8 +503,9 @@ static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___c static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__44; static lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___closed__3; extern lean_object* l_Lean_Meta_instMonadMetaM; -LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at_Lean_Elab_Tactic_Omega_MetaProblem_addFact___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_instantiateMVars___at_Lean_Elab_Tactic_Omega_MetaProblem_addFact___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__10; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___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_Lean_MVarId_withContext___at_Lean_Elab_Tactic_Omega_splitDisjunction___spec__2(lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_trivial___closed__1; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__25; @@ -502,7 +515,7 @@ static lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___closed__1; lean_object* l_Lean_MVarId_assert(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_Omega_cases_u2082___closed__1; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___closed__2; -lean_object* l_Lean_Elab_Tactic_Omega_atomsList___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Elab_Tactic_Omega_atomsList___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_cases_u2082___closed__2; uint8_t l_List_decidableBAll___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__7; @@ -537,15 +550,19 @@ lean_object* l_Lean_MessageData_ofExpr(lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__1___closed__1; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__53; uint8_t l_Lean_LocalDecl_isImplementationDetail(lean_object*); +lean_object* l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___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*); +static lean_object* l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1___closed__9; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__37; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__43; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearCombo___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_Elab_Tactic_Omega_asLinearCombo___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__7___closed__1; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleFinVal___closed__9; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__1___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Omega_evalOmega_declRange___closed__6; -lean_object* l_Lean_Elab_Tactic_Omega_lookup(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Elab_Tactic_Omega_lookup(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___lambda__2___closed__4; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearCombo___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_Omega_MetaProblem_addIntInequality___boxed(lean_object*, 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_Omega_asLinearComboImpl___lambda__8___closed__47; LEAN_EXPORT lean_object* l_Lean_HashMapImp_expand___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__6(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleFinVal___closed__8; @@ -555,7 +572,7 @@ static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___c static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__31; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_processedFacts___default; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__43; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__10; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__41; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___lambda__2___closed__2; @@ -571,9 +588,12 @@ static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__27; static lean_object* l_Lean_Elab_Tactic_Omega_elabOmegaConfig___closed__4; +LEAN_EXPORT lean_object* l_Lean_MVarId_withContext___at_Lean_Elab_Tactic_Omega_splitDisjunction___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*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__4___closed__1; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__2___boxed(lean_object**); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__20___closed__2; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__70; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___boxed(lean_object*, lean_object*, lean_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_abs(lean_object*); uint8_t l_Lean_Elab_Tactic_Omega_Problem_isEmpty(lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__14; @@ -585,9 +605,10 @@ static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___lambda_ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__21(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_Omega_MetaProblem_pushNot___lambda__7(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_Omega_asLinearComboImpl___lambda__4___closed__4; +LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__3___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__11; lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___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_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__3(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_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__41; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__72; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__3___boxed(lean_object**); @@ -602,11 +623,12 @@ static lean_object* l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__23; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__19; lean_object* lean_nat_mod(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(lean_object*, 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_Omega_MetaProblem_addFact(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__65; +LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___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_Omega_asLinearComboImpl___lambda__8___closed__5; static lean_object* l_Lean_Elab_Tactic_Omega_elabOmegaConfig___closed__7; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8(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* l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(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*); static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__7___closed__1; @@ -623,19 +645,19 @@ static lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__2___closed__1; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__14___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__19(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_Omega_asLinearComboImpl_handleNatCast___closed__21; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__3(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__22; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Omega_asLinearComboImpl___spec__3(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__52; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__44; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearCombo___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_Omega_asLinearCombo___lambda__4(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*); static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__4; static lean_object* l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1___closed__4; -lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___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_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___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*); uint8_t l_Lean_Expr_isConstOf(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_replace___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__9(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Tactic_Omega_groundInt_x3f(lean_object*); -lean_object* l_Lean_Elab_Tactic_Omega_commitWhen___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Elab_Tactic_Omega_commitWhen___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_Lean_Elab_Tactic_Omega_splitDisjunction___closed__4; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleFinVal___closed__10; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__71; @@ -647,21 +669,21 @@ static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_problem___default___lam uint8_t l_Lean_Syntax_isNone(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Tactic_Omega_elabOmegaConfig___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_Omega_splitDisjunction___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_throwError___at_Lean_Elab_Tactic_Omega_splitDisjunction___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_Omega_MetaProblem_pushNot___lambda__2___closed__3; uint8_t lean_int_dec_lt(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_mkCoordinateEvalAtomsEq___closed__9; lean_object* lean_nat_sub(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__9; static lean_object* l_Lean_Elab_Tactic_Omega_elabOmegaConfig___closed__10; -lean_object* l_Lean_Elab_Tactic_Omega_Problem_solveEqualities(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Elab_Tactic_Omega_Problem_solveEqualities(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__21; lean_object* lean_nat_mul(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Tactic_Omega_elabOmegaConfig___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___lambda__2___closed__5; static lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___closed__7; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleFinVal___closed__11; -LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___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_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_elabOmegaConfig___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Omega_evalOmega_declRange___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__20___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -685,19 +707,22 @@ static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__7___clo static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__1___closed__2; lean_object* l_Lean_Expr_int_x3f(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__17(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_Omega_omegaImpl___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_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__34; lean_object* l_List_reverse___rarg(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality(lean_object*, lean_object*, 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_Omega_MetaProblem_addIntEquality(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*); static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__3___closed__5; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_Omega_evalOmega_declRange___closed__1; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__6; static lean_object* l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__18; static lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__6___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_Omega_elabOmegaConfig___spec__11(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_PersistentArray_mapMAux___at_Lean_Elab_Tactic_Omega_elabOmegaConfig___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_Tactic_Omega_asLinearComboImpl___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_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__5(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*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__3___closed__3; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___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_Tactic_Omega_asLinearComboImpl___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___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_String_intercalate(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__4; static lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__6___closed__1; @@ -708,17 +733,18 @@ static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___c static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__23; lean_object* l_Lean_instantiateMVarsCore(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_evalUnsafe____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_6_(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at_Lean_Elab_Tactic_Omega_omegaImpl___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_MVarId_assign___at_Lean_Elab_Tactic_Omega_omegaImpl___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_Elab_Tactic_Omega_evalOmega___lambda__1___closed__1; static lean_object* l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__12; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleFinVal___closed__2; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__36; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___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*); size_t lean_usize_add(size_t, size_t); static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__14; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__32; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__19; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__5(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*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__44; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__16; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Omega_evalOmega_declRange___closed__7; @@ -727,7 +753,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_evalUnsafe____x40_Lean_Elab_Ta LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_omegaTactic(lean_object*, 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*); -lean_object* l_Lean_Elab_Tactic_Omega_atoms___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Elab_Tactic_Omega_atoms___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_instToExprInt_mkNat(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_mapTR_loop___at_Lean_Elab_Tactic_Omega_Problem_instToStringProblem___spec__4(lean_object*, lean_object*); @@ -735,17 +761,17 @@ static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___clo static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__2; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__21; static lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__6___closed__3; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___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_Omega_splitDisjunction___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*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__47; lean_object* l_Lean_Elab_Tactic_Omega_atomsCoeffs___boxed(lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__14; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(lean_object*, 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_Omega_asLinearComboImpl_rewrite(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__42; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); static lean_object* l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1___closed__6; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__22; 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_Elab_Tactic_Omega_splitDisjunction___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__7(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_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__17; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__48; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__52; @@ -764,7 +790,7 @@ lean_object* lean_string_append(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__4; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__7___closed__2; 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_Elab_Tactic_Omega_splitDisjunction(lean_object*, 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_Omega_splitDisjunction(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__18; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__3; lean_object* l_Lean_HashMap_toList___at_Lean_Elab_Tactic_Omega_Problem_instToStringProblem___spec__1(lean_object*); @@ -783,6 +809,7 @@ static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___ static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__16; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__60; lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___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*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__42; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__23; @@ -795,13 +822,16 @@ static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___c static lean_object* l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__24; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__3; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__19; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___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_Tactic_Omega_splitDisjunction___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_cases_u2082___closed__4; static lean_object* l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__10; lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_withoutModifyingStateWithInfoAndMessagesImpl___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); +static lean_object* l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1___closed__8; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality___spec__2(lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_problem___default___closed__4; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleFinVal___boxed(lean_object*, lean_object*, 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_Omega_asLinearComboImpl___lambda__3___closed__3; lean_object* l_Lean_mkApp5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_Omega_evalOmega___closed__4; @@ -811,7 +841,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__ lean_object* l_List_enumFrom___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__30; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearCombo(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_Omega_asLinearCombo(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__51; static lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_14702____closed__2; lean_object* l_Lean_Expr_nat_x3f(lean_object*); @@ -822,15 +852,16 @@ static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__1___c lean_object* l_Lean_Elab_Tactic_replaceMainGoal(lean_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_Omega_asLinearComboImpl_handleNatCast___closed__11; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___closed__2; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___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_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___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_Elab_Tactic_Omega_asLinearComboImpl___lambda__5___closed__2; lean_object* l_Lean_HashSet_toList___at_Lean_Elab_Tactic_Omega_lookup___spec__10(lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__22; lean_object* l_Lean_Omega_LinearCombo_sub(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__74; LEAN_EXPORT lean_object* l_Lean_HashMapImp_moveEntries___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__7(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality___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_MVarId_withContext___at_Lean_Elab_Tactic_Omega_splitDisjunction___spec__2___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_Omega_MetaProblem_addIntInequality___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_MVarId_withContext___at_Lean_Elab_Tactic_Omega_splitDisjunction___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__5___closed__1; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__63; lean_object* l_Nat_repr(lean_object*); @@ -841,10 +872,11 @@ static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__11___ static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__20; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_whnfR(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___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_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___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*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMapImp_find_x3f___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___boxed(lean_object*, 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_Elab_Tactic_Omega_splitDisjunction___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_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__11___closed__1; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_problem___default___closed__3; @@ -4224,102 +4256,105 @@ x_3 = l_Lean_Expr_const___override(x_2, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_mkEvalRflProof(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_Tactic_Omega_mkEvalRflProof(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_object* x_11, lean_object* x_12) { _start: { -lean_object* x_11; +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); -x_11 = l_Lean_Elab_Tactic_Omega_atomsCoeffs___rarg(x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_11) == 0) +x_13 = l_Lean_Elab_Tactic_Omega_atomsCoeffs___rarg(x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) { -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_20; -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = lean_ctor_get(x_2, 0); -x_15 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__8; -x_16 = lean_int_dec_le(x_15, x_14); -x_17 = lean_ctor_get(x_2, 1); -x_18 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__17; -x_19 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__21; -x_20 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1(x_18, x_19, x_17); -if (x_16 == 0) +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +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_2, 0); +x_17 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__8; +x_18 = lean_int_dec_le(x_17, x_16); +x_19 = lean_ctor_get(x_2, 1); +x_20 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__17; +x_21 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__21; +x_22 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1(x_20, x_21, x_19); +if (x_18 == 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; -x_21 = lean_int_neg(x_14); -x_22 = l_Int_toNat(x_21); -lean_dec(x_21); -x_23 = l_Lean_instToExprInt_mkNat(x_22); -x_24 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__27; -x_25 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__11; -x_26 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__30; -x_27 = l_Lean_mkApp3(x_24, x_25, x_26, x_23); -x_28 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__7; -x_29 = l_Lean_Expr_app___override(x_28, x_27); -x_30 = l_Lean_Expr_app___override(x_29, x_20); -x_31 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__4; -x_32 = l_Lean_mkAppB(x_31, x_30, x_12); -x_33 = l_Lean_Elab_Tactic_Omega_mkEqReflWithExpectedType(x_1, x_32, x_6, x_7, x_8, x_9, x_13); -return x_33; +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; 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_23 = lean_int_neg(x_16); +x_24 = l_Int_toNat(x_23); +lean_dec(x_23); +x_25 = l_Lean_instToExprInt_mkNat(x_24); +x_26 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__27; +x_27 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__11; +x_28 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__30; +x_29 = l_Lean_mkApp3(x_26, x_27, x_28, x_25); +x_30 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__7; +x_31 = l_Lean_Expr_app___override(x_30, x_29); +x_32 = l_Lean_Expr_app___override(x_31, x_22); +x_33 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__4; +x_34 = l_Lean_mkAppB(x_33, x_32, x_14); +x_35 = l_Lean_Elab_Tactic_Omega_mkEqReflWithExpectedType(x_1, x_34, x_8, x_9, x_10, x_11, x_15); +return x_35; } 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; -x_34 = l_Int_toNat(x_14); -x_35 = l_Lean_instToExprInt_mkNat(x_34); -x_36 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__7; -x_37 = l_Lean_Expr_app___override(x_36, x_35); -x_38 = l_Lean_Expr_app___override(x_37, x_20); -x_39 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__4; -x_40 = l_Lean_mkAppB(x_39, x_38, x_12); -x_41 = l_Lean_Elab_Tactic_Omega_mkEqReflWithExpectedType(x_1, x_40, x_6, x_7, x_8, x_9, x_13); -return x_41; +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_36 = l_Int_toNat(x_16); +x_37 = l_Lean_instToExprInt_mkNat(x_36); +x_38 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__7; +x_39 = l_Lean_Expr_app___override(x_38, x_37); +x_40 = l_Lean_Expr_app___override(x_39, x_22); +x_41 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__4; +x_42 = l_Lean_mkAppB(x_41, x_40, x_14); +x_43 = l_Lean_Elab_Tactic_Omega_mkEqReflWithExpectedType(x_1, x_42, x_8, x_9, x_10, x_11, x_15); +return x_43; } } else { -uint8_t x_42; +uint8_t x_44; +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_42 = !lean_is_exclusive(x_11); -if (x_42 == 0) +x_44 = !lean_is_exclusive(x_13); +if (x_44 == 0) { -return x_11; +return x_13; } else { -lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_43 = lean_ctor_get(x_11, 0); -x_44 = lean_ctor_get(x_11, 1); -lean_inc(x_44); -lean_inc(x_43); -lean_dec(x_11); -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_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_13, 0); +x_46 = lean_ctor_get(x_13, 1); +lean_inc(x_46); +lean_inc(x_45); +lean_dec(x_13); +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_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_mkEvalRflProof___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_Tactic_Omega_mkEvalRflProof___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_11; -x_11 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +uint8_t x_13; lean_object* x_14; +x_13 = lean_unbox(x_6); +lean_dec(x_6); +x_14 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof(x_1, x_2, x_3, x_4, x_5, x_13, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_7); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_11; +return x_14; } } static lean_object* _init_l_Lean_Elab_Tactic_Omega_mkCoordinateEvalAtomsEq___closed__1() { @@ -4419,426 +4454,439 @@ x_1 = lean_mk_string_from_bytes("coordinate_eval_", 16); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_mkCoordinateEvalAtomsEq(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_Tactic_Omega_mkCoordinateEvalAtomsEq(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_object* x_11, lean_object* x_12) { _start: { -lean_object* x_11; uint8_t x_12; -x_11 = lean_unsigned_to_nat(10u); -x_12 = lean_nat_dec_lt(x_2, x_11); -if (x_12 == 0) +lean_object* x_13; uint8_t x_14; +x_13 = lean_unsigned_to_nat(10u); +x_14 = lean_nat_dec_lt(x_2, x_13); +if (x_14 == 0) { -lean_object* x_13; +lean_object* x_15; +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_13 = l_Lean_Elab_Tactic_Omega_atomsCoeffs___rarg(x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_13) == 0) +x_15 = l_Lean_Elab_Tactic_Omega_atomsCoeffs___rarg(x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; 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 = l_Lean_mkNatLit(x_2); -x_17 = l_Lean_Elab_Tactic_Omega_mkCoordinateEvalAtomsEq___closed__4; +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_16 = lean_ctor_get(x_15, 0); lean_inc(x_16); -lean_inc(x_14); -x_18 = l_Lean_mkAppB(x_17, x_14, x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = l_Lean_mkNatLit(x_2); +x_19 = l_Lean_Elab_Tactic_Omega_mkCoordinateEvalAtomsEq___closed__4; +lean_inc(x_18); +lean_inc(x_16); +x_20 = l_Lean_mkAppB(x_19, x_16, x_18); +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_19 = l_Lean_Elab_Tactic_Omega_mkEqReflWithExpectedType(x_1, x_18, x_6, x_7, x_8, x_9, x_15); -if (lean_obj_tag(x_19) == 0) +x_21 = l_Lean_Elab_Tactic_Omega_mkEqReflWithExpectedType(x_1, x_20, x_8, x_9, x_10, x_11, x_17); +if (lean_obj_tag(x_21) == 0) { -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_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_Elab_Tactic_Omega_mkCoordinateEvalAtomsEq___closed__7; -x_23 = l_Lean_mkAppB(x_22, x_16, x_14); +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 = l_Lean_Elab_Tactic_Omega_mkCoordinateEvalAtomsEq___closed__7; +x_25 = l_Lean_mkAppB(x_24, x_18, x_16); +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_24 = l_Lean_Meta_mkEqSymm(x_23, x_6, x_7, x_8, x_9, x_21); -if (lean_obj_tag(x_24) == 0) +x_26 = l_Lean_Meta_mkEqSymm(x_25, x_8, x_9, x_10, x_11, x_23); +if (lean_obj_tag(x_26) == 0) { -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_24, 0); -lean_inc(x_25); -x_26 = lean_ctor_get(x_24, 1); -lean_inc(x_26); -lean_dec(x_24); -x_27 = l_Lean_Meta_mkEqTrans(x_20, x_25, x_6, x_7, x_8, x_9, x_26); -return x_27; +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +lean_dec(x_26); +x_29 = l_Lean_Meta_mkEqTrans(x_22, x_27, x_8, x_9, x_10, x_11, x_28); +return x_29; } else { -uint8_t x_28; -lean_dec(x_20); +uint8_t x_30; +lean_dec(x_22); +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_28 = !lean_is_exclusive(x_24); -if (x_28 == 0) +x_30 = !lean_is_exclusive(x_26); +if (x_30 == 0) { -return x_24; +return x_26; } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_24, 0); -x_30 = lean_ctor_get(x_24, 1); -lean_inc(x_30); -lean_inc(x_29); -lean_dec(x_24); -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_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_26, 0); +x_32 = lean_ctor_get(x_26, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_26); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; } } } else { -uint8_t x_32; +uint8_t x_34; +lean_dec(x_18); lean_dec(x_16); -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); -x_32 = !lean_is_exclusive(x_19); -if (x_32 == 0) +x_34 = !lean_is_exclusive(x_21); +if (x_34 == 0) { -return x_19; +return x_21; } 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_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; +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_21, 0); +x_36 = lean_ctor_get(x_21, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_21); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +return x_37; } } } else { -uint8_t x_36; +uint8_t x_38; +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_36 = !lean_is_exclusive(x_13); -if (x_36 == 0) +x_38 = !lean_is_exclusive(x_15); +if (x_38 == 0) { -return x_13; +return x_15; } 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_39; lean_object* x_40; lean_object* x_41; +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_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_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; 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_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; 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_dec(x_1); -x_40 = l_Lean_Elab_Tactic_Omega_atoms___rarg(x_4, x_5, x_6, x_7, x_8, x_9, x_10); -x_41 = lean_ctor_get(x_40, 0); -lean_inc(x_41); -x_42 = lean_ctor_get(x_40, 1); -lean_inc(x_42); -lean_dec(x_40); -x_43 = lean_box(0); -x_44 = lean_unsigned_to_nat(1u); -x_45 = lean_nat_add(x_2, x_44); -x_46 = lean_array_get_size(x_41); -lean_inc(x_45); -lean_inc(x_41); -x_47 = l_Array_toSubarray___rarg(x_41, x_45, x_46); -x_48 = l_Array_ofSubarray___rarg(x_47); -x_49 = lean_array_to_list(lean_box(0), x_48); -x_50 = l_Lean_Elab_Tactic_Omega_mkCoordinateEvalAtomsEq___closed__8; +x_42 = l_Lean_Elab_Tactic_Omega_atoms___rarg(x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +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_box(0); +x_46 = lean_unsigned_to_nat(1u); +x_47 = lean_nat_add(x_2, x_46); +x_48 = lean_array_get_size(x_43); +lean_inc(x_47); +lean_inc(x_43); +x_49 = l_Array_toSubarray___rarg(x_43, x_47, x_48); +x_50 = l_Array_ofSubarray___rarg(x_49); +x_51 = lean_array_to_list(lean_box(0), x_50); +x_52 = l_Lean_Elab_Tactic_Omega_mkCoordinateEvalAtomsEq___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_51 = l_Lean_Meta_mkListLit(x_50, x_49, x_6, x_7, x_8, x_9, x_42); -if (lean_obj_tag(x_51) == 0) +x_53 = l_Lean_Meta_mkListLit(x_52, x_51, x_8, x_9, x_10, x_11, x_44); +if (lean_obj_tag(x_53) == 0) { -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_52 = lean_ctor_get(x_51, 0); -lean_inc(x_52); -x_53 = lean_ctor_get(x_51, 1); -lean_inc(x_53); -lean_dec(x_51); -x_54 = l_Nat_repr(x_2); -x_55 = l_Lean_Elab_Tactic_Omega_mkCoordinateEvalAtomsEq___closed__10; -x_56 = lean_string_append(x_55, x_54); -lean_dec(x_54); -x_57 = l_Lean_Elab_Tactic_Omega_MetaProblem_problem___default___lambda__1___closed__1; -x_58 = lean_string_append(x_56, x_57); -x_59 = l_Lean_Elab_Tactic_Omega_mkCoordinateEvalAtomsEq___closed__9; -x_60 = l_Lean_Name_str___override(x_59, x_58); -x_61 = l_Lean_Expr_const___override(x_60, x_43); -x_62 = lean_unsigned_to_nat(0u); -x_63 = l_Array_toSubarray___rarg(x_41, x_62, x_45); -x_64 = l_Array_ofSubarray___rarg(x_63); -x_65 = lean_array_push(x_64, x_52); -x_66 = l_Lean_mkAppN(x_61, x_65); -x_67 = l_Lean_Meta_mkEqSymm(x_66, x_6, x_7, x_8, x_9, x_53); -return x_67; +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_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 = l_Nat_repr(x_2); +x_57 = l_Lean_Elab_Tactic_Omega_mkCoordinateEvalAtomsEq___closed__10; +x_58 = lean_string_append(x_57, x_56); +lean_dec(x_56); +x_59 = l_Lean_Elab_Tactic_Omega_MetaProblem_problem___default___lambda__1___closed__1; +x_60 = lean_string_append(x_58, x_59); +x_61 = l_Lean_Elab_Tactic_Omega_mkCoordinateEvalAtomsEq___closed__9; +x_62 = l_Lean_Name_str___override(x_61, x_60); +x_63 = l_Lean_Expr_const___override(x_62, x_45); +x_64 = lean_unsigned_to_nat(0u); +x_65 = l_Array_toSubarray___rarg(x_43, x_64, x_47); +x_66 = l_Array_ofSubarray___rarg(x_65); +x_67 = lean_array_push(x_66, x_54); +x_68 = l_Lean_mkAppN(x_63, x_67); +x_69 = l_Lean_Meta_mkEqSymm(x_68, x_8, x_9, x_10, x_11, x_55); +return x_69; } else { -uint8_t x_68; -lean_dec(x_45); -lean_dec(x_41); +uint8_t x_70; +lean_dec(x_47); +lean_dec(x_43); +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_68 = !lean_is_exclusive(x_51); -if (x_68 == 0) +x_70 = !lean_is_exclusive(x_53); +if (x_70 == 0) { -return x_51; +return x_53; } 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_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_object* x_71; lean_object* x_72; lean_object* x_73; +x_71 = lean_ctor_get(x_53, 0); +x_72 = lean_ctor_get(x_53, 1); +lean_inc(x_72); +lean_inc(x_71); +lean_dec(x_53); +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_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_mkCoordinateEvalAtomsEq___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_Tactic_Omega_mkCoordinateEvalAtomsEq___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_11; -x_11 = l_Lean_Elab_Tactic_Omega_mkCoordinateEvalAtomsEq(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +uint8_t x_13; lean_object* x_14; +x_13 = lean_unbox(x_6); +lean_dec(x_6); +x_14 = l_Lean_Elab_Tactic_Omega_mkCoordinateEvalAtomsEq(x_1, x_2, x_3, x_4, x_5, x_13, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_7); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -return x_11; +return x_14; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(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) { _start: { -lean_object* x_10; +lean_object* x_12; lean_inc(x_1); -x_10 = l_Lean_Elab_Tactic_Omega_lookup(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_10) == 0) -{ -uint8_t x_11; -x_11 = !lean_is_exclusive(x_10); -if (x_11 == 0) +x_12 = l_Lean_Elab_Tactic_Omega_lookup(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_12) == 0) { -lean_object* x_12; uint8_t x_13; -x_12 = lean_ctor_get(x_10, 0); +uint8_t x_13; 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_14; uint8_t x_15; x_14 = lean_ctor_get(x_12, 0); -x_15 = lean_ctor_get(x_12, 1); -x_16 = l_Lean_Omega_LinearCombo_coordinate(x_14); -x_17 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_mkCoordinateEvalAtomsEq___boxed), 10, 2); -lean_closure_set(x_17, 0, x_1); -lean_closure_set(x_17, 1, x_14); -if (lean_obj_tag(x_15) == 0) +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) { -lean_object* x_18; lean_object* x_19; -x_18 = l_Lean_Elab_Tactic_Omega_MetaProblem_problem___default___closed__2; -lean_ctor_set(x_12, 1, x_18); -lean_ctor_set(x_12, 0, x_17); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_16); -lean_ctor_set(x_19, 1, x_12); -lean_ctor_set(x_10, 0, x_19); -return x_10; -} -else +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_16 = lean_ctor_get(x_14, 0); +x_17 = lean_ctor_get(x_14, 1); +x_18 = l_Lean_Omega_LinearCombo_coordinate(x_16); +x_19 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_mkCoordinateEvalAtomsEq___boxed), 12, 2); +lean_closure_set(x_19, 0, x_1); +lean_closure_set(x_19, 1, x_16); +if (lean_obj_tag(x_17) == 0) { lean_object* x_20; lean_object* x_21; -x_20 = lean_ctor_get(x_15, 0); -lean_inc(x_20); -lean_dec(x_15); -lean_ctor_set(x_12, 1, x_20); -lean_ctor_set(x_12, 0, x_17); +x_20 = l_Lean_Elab_Tactic_Omega_MetaProblem_problem___default___closed__2; +lean_ctor_set(x_14, 1, x_20); +lean_ctor_set(x_14, 0, x_19); x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_16); -lean_ctor_set(x_21, 1, x_12); -lean_ctor_set(x_10, 0, x_21); -return x_10; -} +lean_ctor_set(x_21, 0, x_18); +lean_ctor_set(x_21, 1, x_14); +lean_ctor_set(x_12, 0, x_21); +return x_12; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_22 = lean_ctor_get(x_12, 0); -x_23 = lean_ctor_get(x_12, 1); -lean_inc(x_23); +lean_object* x_22; lean_object* x_23; +x_22 = lean_ctor_get(x_17, 0); lean_inc(x_22); -lean_dec(x_12); -x_24 = l_Lean_Omega_LinearCombo_coordinate(x_22); -x_25 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_mkCoordinateEvalAtomsEq___boxed), 10, 2); -lean_closure_set(x_25, 0, x_1); -lean_closure_set(x_25, 1, x_22); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = l_Lean_Elab_Tactic_Omega_MetaProblem_problem___default___closed__2; -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_25); -lean_ctor_set(x_27, 1, x_26); -x_28 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_28, 0, x_24); -lean_ctor_set(x_28, 1, x_27); -lean_ctor_set(x_10, 0, x_28); -return x_10; +lean_dec(x_17); +lean_ctor_set(x_14, 1, x_22); +lean_ctor_set(x_14, 0, x_19); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_18); +lean_ctor_set(x_23, 1, x_14); +lean_ctor_set(x_12, 0, x_23); +return x_12; +} } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_23, 0); -lean_inc(x_29); -lean_dec(x_23); +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_24 = lean_ctor_get(x_14, 0); +x_25 = lean_ctor_get(x_14, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_14); +x_26 = l_Lean_Omega_LinearCombo_coordinate(x_24); +x_27 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_mkCoordinateEvalAtomsEq___boxed), 12, 2); +lean_closure_set(x_27, 0, x_1); +lean_closure_set(x_27, 1, x_24); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = l_Lean_Elab_Tactic_Omega_MetaProblem_problem___default___closed__2; +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); x_30 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_30, 0, x_25); +lean_ctor_set(x_30, 0, x_26); lean_ctor_set(x_30, 1, x_29); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_24); -lean_ctor_set(x_31, 1, x_30); -lean_ctor_set(x_10, 0, x_31); -return x_10; +lean_ctor_set(x_12, 0, x_30); +return x_12; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_25, 0); +lean_inc(x_31); +lean_dec(x_25); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_27); +lean_ctor_set(x_32, 1, x_31); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_26); +lean_ctor_set(x_33, 1, x_32); +lean_ctor_set(x_12, 0, x_33); +return x_12; } } } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_32 = lean_ctor_get(x_10, 0); -x_33 = lean_ctor_get(x_10, 1); -lean_inc(x_33); -lean_inc(x_32); -lean_dec(x_10); -x_34 = lean_ctor_get(x_32, 0); -lean_inc(x_34); -x_35 = lean_ctor_get(x_32, 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; +x_34 = lean_ctor_get(x_12, 0); +x_35 = lean_ctor_get(x_12, 1); lean_inc(x_35); -if (lean_is_exclusive(x_32)) { - lean_ctor_release(x_32, 0); - lean_ctor_release(x_32, 1); - x_36 = x_32; +lean_inc(x_34); +lean_dec(x_12); +x_36 = lean_ctor_get(x_34, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_34, 1); +lean_inc(x_37); +if (lean_is_exclusive(x_34)) { + lean_ctor_release(x_34, 0); + lean_ctor_release(x_34, 1); + x_38 = x_34; } else { - lean_dec_ref(x_32); - x_36 = lean_box(0); -} -x_37 = l_Lean_Omega_LinearCombo_coordinate(x_34); -x_38 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_mkCoordinateEvalAtomsEq___boxed), 10, 2); -lean_closure_set(x_38, 0, x_1); -lean_closure_set(x_38, 1, x_34); -if (lean_obj_tag(x_35) == 0) -{ -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_39 = l_Lean_Elab_Tactic_Omega_MetaProblem_problem___default___closed__2; -if (lean_is_scalar(x_36)) { - x_40 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_34); + x_38 = lean_box(0); +} +x_39 = l_Lean_Omega_LinearCombo_coordinate(x_36); +x_40 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_mkCoordinateEvalAtomsEq___boxed), 12, 2); +lean_closure_set(x_40, 0, x_1); +lean_closure_set(x_40, 1, x_36); +if (lean_obj_tag(x_37) == 0) +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_41 = l_Lean_Elab_Tactic_Omega_MetaProblem_problem___default___closed__2; +if (lean_is_scalar(x_38)) { + x_42 = lean_alloc_ctor(0, 2, 0); } else { - x_40 = x_36; + x_42 = x_38; } -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -x_41 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_41, 0, x_37); -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_33); -return x_42; +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_41); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_39); +lean_ctor_set(x_43, 1, x_42); +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_44, 1, x_35); +return x_44; } else { -lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_43 = lean_ctor_get(x_35, 0); -lean_inc(x_43); -lean_dec(x_35); -if (lean_is_scalar(x_36)) { - x_44 = lean_alloc_ctor(0, 2, 0); +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_45 = lean_ctor_get(x_37, 0); +lean_inc(x_45); +lean_dec(x_37); +if (lean_is_scalar(x_38)) { + x_46 = lean_alloc_ctor(0, 2, 0); } else { - x_44 = x_36; + x_46 = x_38; } -lean_ctor_set(x_44, 0, x_38); -lean_ctor_set(x_44, 1, x_43); -x_45 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_45, 0, x_37); -lean_ctor_set(x_45, 1, 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_33); -return x_46; +lean_ctor_set(x_46, 0, x_40); +lean_ctor_set(x_46, 1, x_45); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_39); +lean_ctor_set(x_47, 1, x_46); +x_48 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_35); +return x_48; } } } else { -uint8_t x_47; +uint8_t x_49; lean_dec(x_1); -x_47 = !lean_is_exclusive(x_10); -if (x_47 == 0) +x_49 = !lean_is_exclusive(x_12); +if (x_49 == 0) { -return x_10; +return x_12; } else { -lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_48 = lean_ctor_get(x_10, 0); -x_49 = lean_ctor_get(x_10, 1); -lean_inc(x_49); -lean_inc(x_48); -lean_dec(x_10); -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_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_12, 0); +x_51 = lean_ctor_get(x_12, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_12); +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_Tactic_Omega_mkAtomLinearCombo___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_5); +lean_dec(x_5); +x_13 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_2, x_3, x_4, x_12, x_6, x_7, x_8, x_9, x_10, x_11); +return x_13; +} } static lean_object* _init_l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1___closed__1() { _start: @@ -4870,18 +4918,36 @@ return x_2; static lean_object* _init_l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1___closed__4() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; +lean_object* x_1; lean_object* x_2; x_1 = l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1___closed__3; +x_2 = l_ReaderT_instMonadReaderT___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1___closed__4; +x_2 = l_ReaderT_instMonadReaderT___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1___closed__5; x_2 = l_Lean_instInhabitedExpr; x_3 = l_instInhabited___rarg(x_1, x_2); return x_3; } } -static lean_object* _init_l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1___closed__5() { +static lean_object* _init_l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1___closed__4; +x_1 = l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1___closed__6; x_2 = l_Lean_Elab_Tactic_Omega_MetaProblem_problem___default___closed__2; x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -4889,44 +4955,45 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1___closed__6() { +static lean_object* _init_l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Omega_LinearCombo_instInhabitedLinearCombo; -x_2 = l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1___closed__5; +x_2 = l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1___closed__7; x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1___closed__7() { +static lean_object* _init_l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1___closed__3; -x_2 = l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1___closed__6; +x_1 = l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1___closed__5; +x_2 = l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1___closed__8; x_3 = l_instInhabited___rarg(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___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_EXPORT lean_object* l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1(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) { _start: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1___closed__7; -x_11 = lean_panic_fn(x_10, x_1); -x_12 = lean_apply_8(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_12; +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_12 = l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1___closed__9; +x_13 = lean_panic_fn(x_12, x_1); +x_14 = lean_box(x_5); +x_15 = lean_apply_10(x_13, x_2, x_3, x_4, x_14, x_6, x_7, x_8, x_9, x_10, x_11); +return x_15; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___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_Elab_Tactic_Omega_asLinearComboImpl_rewrite___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_object* x_11, lean_object* x_12) { _start: { -lean_object* x_11; -x_11 = l_Lean_Meta_mkEqTrans(x_1, x_2, x_6, x_7, x_8, x_9, x_10); -return x_11; +lean_object* x_13; +x_13 = l_Lean_Meta_mkEqTrans(x_1, x_2, x_8, x_9, x_10, x_11, x_12); +return x_13; } } static lean_object* _init_l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___lambda__2___closed__1() { @@ -4984,238 +5051,239 @@ 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_Tactic_Omega_asLinearComboImpl_rewrite___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_Elab_Tactic_Omega_asLinearComboImpl_rewrite___lambda__2(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_object* x_11, lean_object* x_12) { _start: { -lean_object* x_11; +lean_object* x_13; lean_dec(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_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) +x_13 = lean_infer_type(x_1, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___lambda__2___closed__2; -x_15 = lean_unsigned_to_nat(3u); -x_16 = l_Lean_Expr_isAppOfArity(x_12, x_14, x_15); -if (x_16 == 0) +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___lambda__2___closed__2; +x_17 = lean_unsigned_to_nat(3u); +x_18 = l_Lean_Expr_isAppOfArity(x_14, x_16, x_17); +if (x_18 == 0) { -lean_object* x_17; lean_object* x_18; -lean_dec(x_12); +lean_object* x_19; lean_object* x_20; +lean_dec(x_14); lean_dec(x_1); -x_17 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___lambda__2___closed__6; -x_18 = l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1(x_17, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); -return x_18; +x_19 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___lambda__2___closed__6; +x_20 = l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1(x_19, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_15); +return x_20; } else { -lean_object* x_19; lean_object* x_20; -x_19 = l_Lean_Expr_appArg_x21(x_12); -lean_dec(x_12); -x_20 = l_Lean_Elab_Tactic_Omega_asLinearCombo(x_19, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); -if (lean_obj_tag(x_20) == 0) -{ -lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_21, 1); -lean_inc(x_22); -x_23 = !lean_is_exclusive(x_20); -if (x_23 == 0) +lean_object* x_21; lean_object* x_22; +x_21 = l_Lean_Expr_appArg_x21(x_14); +lean_dec(x_14); +x_22 = l_Lean_Elab_Tactic_Omega_asLinearCombo(x_21, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_15); +if (lean_obj_tag(x_22) == 0) { -lean_object* x_24; uint8_t x_25; -x_24 = lean_ctor_get(x_20, 0); -lean_dec(x_24); -x_25 = !lean_is_exclusive(x_21); +lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +x_25 = !lean_is_exclusive(x_22); if (x_25 == 0) { lean_object* x_26; uint8_t x_27; -x_26 = lean_ctor_get(x_21, 1); +x_26 = lean_ctor_get(x_22, 0); lean_dec(x_26); -x_27 = !lean_is_exclusive(x_22); +x_27 = !lean_is_exclusive(x_23); if (x_27 == 0) { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_22, 0); -x_29 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___lambda__1___boxed), 10, 1); -lean_closure_set(x_29, 0, x_1); -x_30 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg), 10, 2); -lean_closure_set(x_30, 0, x_28); -lean_closure_set(x_30, 1, x_29); -lean_ctor_set(x_22, 0, x_30); -return x_20; +lean_object* x_28; uint8_t x_29; +x_28 = lean_ctor_get(x_23, 1); +lean_dec(x_28); +x_29 = !lean_is_exclusive(x_24); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_24, 0); +x_31 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___lambda__1___boxed), 12, 1); +lean_closure_set(x_31, 0, x_1); +x_32 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg___boxed), 12, 2); +lean_closure_set(x_32, 0, x_30); +lean_closure_set(x_32, 1, x_31); +lean_ctor_set(x_24, 0, x_32); +return x_22; } else { -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_22, 0); -x_32 = lean_ctor_get(x_22, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_22); -x_33 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___lambda__1___boxed), 10, 1); -lean_closure_set(x_33, 0, x_1); -x_34 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg), 10, 2); -lean_closure_set(x_34, 0, x_31); -lean_closure_set(x_34, 1, 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); -lean_ctor_set(x_21, 1, x_35); -return x_20; +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_33 = lean_ctor_get(x_24, 0); +x_34 = lean_ctor_get(x_24, 1); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_24); +x_35 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___lambda__1___boxed), 12, 1); +lean_closure_set(x_35, 0, x_1); +x_36 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg___boxed), 12, 2); +lean_closure_set(x_36, 0, x_33); +lean_closure_set(x_36, 1, 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_34); +lean_ctor_set(x_23, 1, x_37); +return x_22; } } else { -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_36 = lean_ctor_get(x_21, 0); -lean_inc(x_36); -lean_dec(x_21); -x_37 = lean_ctor_get(x_22, 0); -lean_inc(x_37); -x_38 = lean_ctor_get(x_22, 1); +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_38 = lean_ctor_get(x_23, 0); lean_inc(x_38); -if (lean_is_exclusive(x_22)) { - lean_ctor_release(x_22, 0); - lean_ctor_release(x_22, 1); - x_39 = x_22; +lean_dec(x_23); +x_39 = lean_ctor_get(x_24, 0); +lean_inc(x_39); +x_40 = lean_ctor_get(x_24, 1); +lean_inc(x_40); +if (lean_is_exclusive(x_24)) { + lean_ctor_release(x_24, 0); + lean_ctor_release(x_24, 1); + x_41 = x_24; } else { - lean_dec_ref(x_22); - x_39 = lean_box(0); -} -x_40 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___lambda__1___boxed), 10, 1); -lean_closure_set(x_40, 0, x_1); -x_41 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg), 10, 2); -lean_closure_set(x_41, 0, x_37); -lean_closure_set(x_41, 1, x_40); -if (lean_is_scalar(x_39)) { - x_42 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_24); + x_41 = lean_box(0); +} +x_42 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___lambda__1___boxed), 12, 1); +lean_closure_set(x_42, 0, x_1); +x_43 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg___boxed), 12, 2); +lean_closure_set(x_43, 0, x_39); +lean_closure_set(x_43, 1, x_42); +if (lean_is_scalar(x_41)) { + x_44 = lean_alloc_ctor(0, 2, 0); } else { - x_42 = x_39; + x_44 = x_41; } -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_38); -x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_36); -lean_ctor_set(x_43, 1, x_42); -lean_ctor_set(x_20, 0, x_43); -return x_20; +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_44, 1, x_40); +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_38); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_22, 0, x_45); +return x_22; } } 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; -x_44 = lean_ctor_get(x_20, 1); -lean_inc(x_44); -lean_dec(x_20); -x_45 = lean_ctor_get(x_21, 0); -lean_inc(x_45); -if (lean_is_exclusive(x_21)) { - lean_ctor_release(x_21, 0); - lean_ctor_release(x_21, 1); - x_46 = x_21; +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_46 = lean_ctor_get(x_22, 1); +lean_inc(x_46); +lean_dec(x_22); +x_47 = lean_ctor_get(x_23, 0); +lean_inc(x_47); +if (lean_is_exclusive(x_23)) { + lean_ctor_release(x_23, 0); + lean_ctor_release(x_23, 1); + x_48 = x_23; } else { - lean_dec_ref(x_21); - x_46 = lean_box(0); + lean_dec_ref(x_23); + x_48 = lean_box(0); } -x_47 = lean_ctor_get(x_22, 0); -lean_inc(x_47); -x_48 = lean_ctor_get(x_22, 1); -lean_inc(x_48); -if (lean_is_exclusive(x_22)) { - lean_ctor_release(x_22, 0); - lean_ctor_release(x_22, 1); - x_49 = x_22; +x_49 = lean_ctor_get(x_24, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_24, 1); +lean_inc(x_50); +if (lean_is_exclusive(x_24)) { + lean_ctor_release(x_24, 0); + lean_ctor_release(x_24, 1); + x_51 = x_24; } else { - lean_dec_ref(x_22); - x_49 = lean_box(0); -} -x_50 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___lambda__1___boxed), 10, 1); -lean_closure_set(x_50, 0, x_1); -x_51 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg), 10, 2); -lean_closure_set(x_51, 0, x_47); -lean_closure_set(x_51, 1, x_50); -if (lean_is_scalar(x_49)) { - x_52 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_24); + x_51 = lean_box(0); +} +x_52 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___lambda__1___boxed), 12, 1); +lean_closure_set(x_52, 0, x_1); +x_53 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg___boxed), 12, 2); +lean_closure_set(x_53, 0, x_49); +lean_closure_set(x_53, 1, x_52); +if (lean_is_scalar(x_51)) { + x_54 = lean_alloc_ctor(0, 2, 0); } else { - x_52 = x_49; + x_54 = x_51; } -lean_ctor_set(x_52, 0, x_51); -lean_ctor_set(x_52, 1, x_48); -if (lean_is_scalar(x_46)) { - x_53 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_53); +lean_ctor_set(x_54, 1, x_50); +if (lean_is_scalar(x_48)) { + x_55 = lean_alloc_ctor(0, 2, 0); } else { - x_53 = x_46; + x_55 = x_48; } -lean_ctor_set(x_53, 0, x_45); -lean_ctor_set(x_53, 1, 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_44); -return x_54; +lean_ctor_set(x_55, 0, x_47); +lean_ctor_set(x_55, 1, x_54); +x_56 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_46); +return x_56; } } else { -uint8_t x_55; +uint8_t x_57; lean_dec(x_1); -x_55 = !lean_is_exclusive(x_20); -if (x_55 == 0) +x_57 = !lean_is_exclusive(x_22); +if (x_57 == 0) { -return x_20; +return x_22; } else { -lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_56 = lean_ctor_get(x_20, 0); -x_57 = lean_ctor_get(x_20, 1); -lean_inc(x_57); -lean_inc(x_56); -lean_dec(x_20); -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; +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 { -uint8_t x_59; +uint8_t x_61; +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_1); -x_59 = !lean_is_exclusive(x_11); -if (x_59 == 0) +x_61 = !lean_is_exclusive(x_13); +if (x_61 == 0) { -return x_11; +return x_13; } else { -lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_11, 0); -x_61 = lean_ctor_get(x_11, 1); -lean_inc(x_61); -lean_inc(x_60); -lean_dec(x_11); -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_62; lean_object* x_63; lean_object* x_64; +x_62 = lean_ctor_get(x_13, 0); +x_63 = lean_ctor_get(x_13, 1); +lean_inc(x_63); +lean_inc(x_62); +lean_dec(x_13); +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; } } } @@ -5298,111 +5366,112 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_Tactic_Omega_asLinearComboImpl_rewrite(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_object* x_11, lean_object* x_12) { _start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_11 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__2; -x_12 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9, 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_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_13 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__2; +x_14 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_13, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_unbox(x_15); +lean_dec(x_15); +if (x_16 == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; +lean_object* x_17; lean_object* x_18; lean_object* x_19; 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_Omega_asLinearComboImpl_rewrite___lambda__2(x_2, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_15); -return x_17; +x_17 = lean_ctor_get(x_14, 1); +lean_inc(x_17); +lean_dec(x_14); +x_18 = lean_box(0); +x_19 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___lambda__2(x_2, x_18, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_17); +return x_19; } else { -lean_object* x_18; lean_object* x_19; -x_18 = lean_ctor_get(x_12, 1); -lean_inc(x_18); -lean_dec(x_12); +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_14, 1); +lean_inc(x_20); +lean_dec(x_14); +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); lean_inc(x_2); -x_19 = lean_infer_type(x_2, x_6, x_7, x_8, x_9, x_18); -if (lean_obj_tag(x_19) == 0) +x_21 = lean_infer_type(x_2, x_8, x_9, x_10, x_11, x_20); +if (lean_obj_tag(x_21) == 0) { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -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_MessageData_ofExpr(x_1); -x_23 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___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_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__6; +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; 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_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_MessageData_ofExpr(x_1); +x_25 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___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_inc(x_2); -x_27 = l_Lean_MessageData_ofExpr(x_2); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_24); +x_27 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__6; 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_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__8; +lean_inc(x_2); +x_29 = l_Lean_MessageData_ofExpr(x_2); 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_MessageData_ofExpr(x_20); +x_31 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__8; 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 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__9; +x_33 = l_Lean_MessageData_ofExpr(x_22); 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_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_11, x_34, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_21); -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_Elab_Tactic_Omega_asLinearComboImpl_rewrite___lambda__2(x_2, x_36, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_37); -return x_38; +x_35 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__9; +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_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_13, x_36, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_23); +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 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___lambda__2(x_2, x_38, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_39); +return x_40; } else { -uint8_t x_39; +uint8_t x_41; +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_39 = !lean_is_exclusive(x_19); -if (x_39 == 0) +x_41 = !lean_is_exclusive(x_21); +if (x_41 == 0) { -return x_19; +return x_21; } else { -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_19, 0); -x_41 = lean_ctor_get(x_19, 1); -lean_inc(x_41); -lean_inc(x_40); -lean_dec(x_19); -x_42 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_42, 0, x_40); -lean_ctor_set(x_42, 1, x_41); -return x_42; +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_21, 0); +x_43 = lean_ctor_get(x_21, 1); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_21); +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; } } } @@ -5458,55 +5527,59 @@ lean_dec(x_2); return x_8; } } -LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_asLinearCombo___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_object* x_9) { +LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__3___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_10; +lean_object* x_12; lean_object* x_13; +x_12 = lean_box(x_5); +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); -x_10 = lean_apply_7(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_10) == 0) +x_13 = lean_apply_9(x_1, x_3, x_4, x_12, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_13) == 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 = lean_apply_8(x_2, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_12); -return x_13; +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = lean_box(x_5); +x_17 = lean_apply_10(x_2, x_14, x_3, x_4, x_16, x_6, x_7, x_8, x_9, x_10, x_15); +return x_17; } else { -uint8_t x_14; +uint8_t 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_4); lean_dec(x_3); lean_dec(x_2); -x_14 = !lean_is_exclusive(x_10); -if (x_14 == 0) +x_18 = !lean_is_exclusive(x_13); +if (x_18 == 0) { -return x_10; +return x_13; } else { -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_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; } } } @@ -5515,7 +5588,7 @@ LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_asLinearComb _start: { lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__3___rarg), 9, 0); +x_3 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__3___rarg___boxed), 11, 0); return x_3; } } @@ -5835,136 +5908,137 @@ return x_38; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearCombo___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_Tactic_Omega_asLinearCombo___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) { _start: { -lean_object* x_8; uint8_t x_9; -x_8 = lean_st_mk_ref(x_1, x_7); -x_9 = !lean_is_exclusive(x_8); -if (x_9 == 0) +lean_object* x_10; uint8_t x_11; +x_10 = lean_st_mk_ref(x_1, x_9); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) { -return x_8; +return x_10; } 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_Tactic_Omega_asLinearCombo___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_inc(x_2); -x_10 = lean_apply_8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_10) == 0) -{ -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_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_10, 0); +x_13 = lean_ctor_get(x_10, 1); +lean_inc(x_13); lean_inc(x_12); lean_dec(x_10); -x_13 = lean_st_ref_get(x_2, x_12); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_13); +return x_14; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearCombo___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, 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 = lean_box(x_5); +lean_inc(x_2); +x_13 = lean_apply_10(x_1, x_2, x_3, x_4, x_12, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = lean_st_ref_get(x_2, x_15); lean_dec(x_2); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) { -lean_object* x_15; lean_object* x_16; -x_15 = lean_ctor_get(x_13, 0); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_11); -lean_ctor_set(x_16, 1, x_15); -lean_ctor_set(x_13, 0, x_16); -return x_13; +lean_object* x_18; lean_object* x_19; +x_18 = lean_ctor_get(x_16, 0); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_14); +lean_ctor_set(x_19, 1, x_18); +lean_ctor_set(x_16, 0, x_19); +return x_16; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_17 = lean_ctor_get(x_13, 0); -x_18 = lean_ctor_get(x_13, 1); -lean_inc(x_18); -lean_inc(x_17); -lean_dec(x_13); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_11); -lean_ctor_set(x_19, 1, x_17); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_18); -return x_20; +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(0, 2, 0); +lean_ctor_set(x_22, 0, x_14); +lean_ctor_set(x_22, 1, 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_21; +uint8_t x_24; lean_dec(x_2); -x_21 = !lean_is_exclusive(x_10); -if (x_21 == 0) +x_24 = !lean_is_exclusive(x_13); +if (x_24 == 0) { -return x_10; +return x_13; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_10, 0); -x_23 = lean_ctor_get(x_10, 1); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_10); -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_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_13, 0); +x_26 = lean_ctor_get(x_13, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_13); +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_Elab_Tactic_Omega_asLinearCombo___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_Tactic_Omega_asLinearCombo___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) { _start: { -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_8); -return x_10; +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_10); +return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearCombo___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_Elab_Tactic_Omega_asLinearCombo___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, 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) { _start: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_12 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); -lean_closure_set(x_12, 0, x_1); -x_13 = l_Lean_Elab_Tactic_Omega_MetaProblem_problem___default___closed__2; -x_14 = lean_alloc_ctor(0, 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_2); -lean_ctor_set(x_15, 1, x_14); +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_14 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); +lean_closure_set(x_14, 0, x_1); +x_15 = l_Lean_Elab_Tactic_Omega_MetaProblem_problem___default___closed__2; x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_11); -return x_16; +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_2); +lean_ctor_set(x_17, 1, 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_13); +return x_18; } } static lean_object* _init_l_Lean_Elab_Tactic_Omega_asLinearCombo___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__3___boxed), 8, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__3___boxed), 10, 0); return x_1; } } @@ -5985,177 +6059,179 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearCombo(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearCombo(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) { _start: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_10 = lean_st_ref_get(x_2, x_9); -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_10, 1); -lean_inc(x_12); -lean_dec(x_10); +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_12 = lean_st_ref_get(x_2, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); lean_inc(x_1); -x_13 = l_Lean_HashMapImp_find_x3f___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__1(x_11, x_1); -if (lean_obj_tag(x_13) == 0) +x_15 = l_Lean_HashMapImp_find_x3f___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__1(x_13, x_1); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_14; +lean_object* x_16; lean_inc(x_2); lean_inc(x_1); -x_14 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_12); -if (lean_obj_tag(x_14) == 0) +x_16 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; -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_st_ref_take(x_2, x_16); -x_18 = lean_ctor_get(x_17, 0); +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); lean_inc(x_18); -x_19 = lean_ctor_get(x_17, 1); -lean_inc(x_19); -lean_dec(x_17); -x_20 = lean_ctor_get(x_15, 0); +lean_dec(x_16); +x_19 = lean_st_ref_take(x_2, x_18); +x_20 = lean_ctor_get(x_19, 0); lean_inc(x_20); -x_21 = lean_ctor_get(x_15, 1); +x_21 = lean_ctor_get(x_19, 1); lean_inc(x_21); -x_22 = lean_ctor_get(x_21, 0); +lean_dec(x_19); +x_22 = lean_ctor_get(x_17, 0); lean_inc(x_22); -lean_dec(x_21); -lean_inc(x_18); -x_23 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__1___boxed), 7, 1); -lean_closure_set(x_23, 0, x_18); -x_24 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); -lean_closure_set(x_24, 0, x_23); -x_25 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__2), 9, 1); -lean_closure_set(x_25, 0, x_22); -x_26 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__3___rarg), 9, 2); -lean_closure_set(x_26, 0, x_24); -lean_closure_set(x_26, 1, x_25); -x_27 = l_Lean_Elab_Tactic_Omega_asLinearCombo___closed__1; -x_28 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__3___rarg), 9, 2); +x_23 = lean_ctor_get(x_17, 1); +lean_inc(x_23); +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +lean_dec(x_23); +lean_inc(x_20); +x_25 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__1___boxed), 9, 1); +lean_closure_set(x_25, 0, x_20); +x_26 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); +lean_closure_set(x_26, 0, x_25); +x_27 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__2___boxed), 11, 1); +lean_closure_set(x_27, 0, x_24); +x_28 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__3___rarg___boxed), 11, 2); lean_closure_set(x_28, 0, x_26); lean_closure_set(x_28, 1, x_27); -x_29 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_29, 0, x_20); -lean_ctor_set(x_29, 1, x_28); -x_30 = l_Lean_HashMap_insert___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__4(x_18, x_1, x_29); -x_31 = lean_st_ref_set(x_2, x_30, x_19); +x_29 = l_Lean_Elab_Tactic_Omega_asLinearCombo___closed__1; +x_30 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__3___rarg___boxed), 11, 2); +lean_closure_set(x_30, 0, x_28); +lean_closure_set(x_30, 1, x_29); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_22); +lean_ctor_set(x_31, 1, x_30); +x_32 = l_Lean_HashMap_insert___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__4(x_20, x_1, x_31); +x_33 = lean_st_ref_set(x_2, x_32, x_21); lean_dec(x_2); -x_32 = !lean_is_exclusive(x_31); -if (x_32 == 0) +x_34 = !lean_is_exclusive(x_33); +if (x_34 == 0) { -lean_object* x_33; -x_33 = lean_ctor_get(x_31, 0); -lean_dec(x_33); -lean_ctor_set(x_31, 0, x_15); -return x_31; +lean_object* x_35; +x_35 = lean_ctor_get(x_33, 0); +lean_dec(x_35); +lean_ctor_set(x_33, 0, x_17); +return x_33; } else { -lean_object* x_34; lean_object* x_35; -x_34 = lean_ctor_get(x_31, 1); -lean_inc(x_34); -lean_dec(x_31); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_15); -lean_ctor_set(x_35, 1, x_34); -return x_35; +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_17); +lean_ctor_set(x_37, 1, x_36); +return x_37; } } else { -uint8_t x_36; +uint8_t x_38; lean_dec(x_2); lean_dec(x_1); -x_36 = !lean_is_exclusive(x_14); -if (x_36 == 0) +x_38 = !lean_is_exclusive(x_16); +if (x_38 == 0) { -return x_14; +return x_16; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_14, 0); -x_38 = lean_ctor_get(x_14, 1); -lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_14); -x_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_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; } } } 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; uint8_t x_46; -x_40 = lean_ctor_get(x_13, 0); -lean_inc(x_40); -lean_dec(x_13); -x_41 = lean_ctor_get(x_40, 0); -lean_inc(x_41); -x_42 = lean_ctor_get(x_40, 1); +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; +x_42 = lean_ctor_get(x_15, 0); lean_inc(x_42); -lean_dec(x_40); -x_43 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__2; -x_44 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_43, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_12); -x_45 = lean_ctor_get(x_44, 0); -lean_inc(x_45); -x_46 = lean_unbox(x_45); -lean_dec(x_45); -if (x_46 == 0) +lean_dec(x_15); +x_43 = lean_ctor_get(x_42, 0); +lean_inc(x_43); +x_44 = lean_ctor_get(x_42, 1); +lean_inc(x_44); +lean_dec(x_42); +x_45 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__2; +x_46 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_45, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +x_47 = lean_ctor_get(x_46, 0); +lean_inc(x_47); +x_48 = lean_unbox(x_47); +lean_dec(x_47); +if (x_48 == 0) { -lean_object* x_47; lean_object* x_48; lean_object* x_49; +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_dec(x_1); -x_47 = lean_ctor_get(x_44, 1); -lean_inc(x_47); -lean_dec(x_44); -x_48 = lean_box(0); -x_49 = l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__4(x_42, x_41, x_48, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_47); +x_49 = lean_ctor_get(x_46, 1); +lean_inc(x_49); +lean_dec(x_46); +x_50 = lean_box(0); +x_51 = l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__4(x_44, x_43, x_50, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_49); +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_49; +return x_51; } 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; -x_50 = lean_ctor_get(x_44, 1); -lean_inc(x_50); -lean_dec(x_44); -x_51 = l_Lean_MessageData_ofExpr(x_1); -x_52 = l_Lean_Elab_Tactic_Omega_asLinearCombo___closed__3; -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_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__9; +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 = lean_ctor_get(x_46, 1); +lean_inc(x_52); +lean_dec(x_46); +x_53 = l_Lean_MessageData_ofExpr(x_1); +x_54 = l_Lean_Elab_Tactic_Omega_asLinearCombo___closed__3; 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 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_43, x_55, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_50); -x_57 = lean_ctor_get(x_56, 0); -lean_inc(x_57); -x_58 = lean_ctor_get(x_56, 1); -lean_inc(x_58); -lean_dec(x_56); -x_59 = l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__4(x_42, x_41, x_57, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_58); +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_53); +x_56 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__9; +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_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_45, x_57, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_52); +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_Tactic_Omega_asLinearCombo___lambda__4(x_44, x_43, x_59, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_60); +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_57); -return x_59; +lean_dec(x_59); +return x_61; } } } @@ -6221,12 +6297,12 @@ x_3 = lean_int_dec_eq(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__2(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_8; -x_8 = l_Lean_Meta_mkDecideProof(x_1, x_3, x_4, x_5, x_6, x_7); -return x_8; +lean_object* x_10; +x_10 = l_Lean_Meta_mkDecideProof(x_1, x_5, x_6, x_7, x_8, x_9); +return x_10; } } static lean_object* _init_l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__3___closed__1() { @@ -6278,220 +6354,195 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___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, uint8_t x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19) { _start: { -lean_object* x_18; +lean_object* x_20; +lean_inc(x_18); +lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -x_18 = l_Lean_Elab_Tactic_Omega_atomsCoeffs___rarg(x_11, x_12, x_13, x_14, x_15, x_16, x_17); -if (lean_obj_tag(x_18) == 0) +x_20 = l_Lean_Elab_Tactic_Omega_atomsCoeffs___rarg(x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); +if (lean_obj_tag(x_20) == 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; lean_object* x_28; uint8_t x_29; lean_object* x_30; lean_object* x_31; lean_object* x_78; -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_Elab_Tactic_Omega_asLinearComboImpl___lambda__3___closed__2; -lean_inc(x_1); -x_22 = l_Lean_Expr_const___override(x_21, x_1); -x_23 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__6; +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_82; +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_Tactic_Omega_asLinearComboImpl___lambda__3___closed__2; lean_inc(x_1); x_24 = l_Lean_Expr_const___override(x_23, x_1); -x_25 = lean_ctor_get(x_2, 0); -x_26 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__8; -x_27 = lean_int_dec_le(x_26, x_25); -x_28 = lean_ctor_get(x_3, 0); -x_29 = lean_int_dec_le(x_26, x_28); +x_25 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__6; +lean_inc(x_1); +x_26 = l_Lean_Expr_const___override(x_25, x_1); +x_27 = lean_ctor_get(x_2, 0); +x_28 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__8; +x_29 = lean_int_dec_le(x_28, x_27); +x_30 = lean_ctor_get(x_3, 0); +x_31 = lean_int_dec_le(x_28, x_30); +x_32 = lean_box(x_13); +lean_inc(x_18); +lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); -lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); -x_30 = lean_apply_8(x_4, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_20); -if (x_27 == 0) +x_33 = lean_apply_10(x_4, x_10, x_11, x_12, x_32, x_14, x_15, x_16, x_17, x_18, x_22); +if (x_29 == 0) { -lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; -x_117 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__25; +lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; +x_122 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__25; lean_inc(x_1); -x_118 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_118, 0, x_117); -lean_ctor_set(x_118, 1, x_1); -x_119 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__24; -x_120 = l_Lean_Expr_const___override(x_119, x_118); -x_121 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__29; +x_123 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_123, 0, x_122); +lean_ctor_set(x_123, 1, x_1); +x_124 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__24; +x_125 = l_Lean_Expr_const___override(x_124, x_123); +x_126 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__29; lean_inc(x_1); -x_122 = l_Lean_Expr_const___override(x_121, x_1); -x_123 = lean_int_neg(x_25); -x_124 = l_Int_toNat(x_123); -lean_dec(x_123); -x_125 = l_Lean_instToExprInt_mkNat(x_124); +x_127 = l_Lean_Expr_const___override(x_126, x_1); +x_128 = lean_int_neg(x_27); +x_129 = l_Int_toNat(x_128); +lean_dec(x_128); +x_130 = l_Lean_instToExprInt_mkNat(x_129); lean_inc(x_5); -x_126 = l_Lean_mkApp3(x_120, x_5, x_122, x_125); -lean_inc(x_24); -x_127 = l_Lean_Expr_app___override(x_24, x_126); -x_128 = l_Lean_Expr_app___override(x_127, x_8); -if (x_29 == 0) +x_131 = l_Lean_mkApp3(x_125, x_5, x_127, x_130); +lean_inc(x_26); +x_132 = l_Lean_Expr_app___override(x_26, x_131); +x_133 = l_Lean_Expr_app___override(x_132, x_8); +if (x_31 == 0) { -x_31 = x_128; -goto block_77; +x_34 = x_133; +goto block_81; } else { lean_dec(x_5); lean_dec(x_1); -x_78 = x_128; -goto block_116; +x_82 = x_133; +goto block_121; } } else { -lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; -x_129 = l_Int_toNat(x_25); -x_130 = l_Lean_instToExprInt_mkNat(x_129); -lean_inc(x_24); -x_131 = l_Lean_Expr_app___override(x_24, x_130); -x_132 = l_Lean_Expr_app___override(x_131, x_8); -if (x_29 == 0) +lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; +x_134 = l_Int_toNat(x_27); +x_135 = l_Lean_instToExprInt_mkNat(x_134); +lean_inc(x_26); +x_136 = l_Lean_Expr_app___override(x_26, x_135); +x_137 = l_Lean_Expr_app___override(x_136, x_8); +if (x_31 == 0) { -x_31 = x_132; -goto block_77; +x_34 = x_137; +goto block_81; } else { lean_dec(x_5); lean_dec(x_1); -x_78 = x_132; -goto block_116; +x_82 = x_137; +goto block_121; } } -block_77: +block_81: { -if (lean_obj_tag(x_30) == 0) +if (lean_obj_tag(x_33) == 0) { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_32 = lean_ctor_get(x_30, 0); -lean_inc(x_32); -x_33 = lean_ctor_get(x_30, 1); -lean_inc(x_33); -lean_dec(x_30); -x_34 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__25; +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; 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_35 = lean_ctor_get(x_33, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_33, 1); +lean_inc(x_36); +lean_dec(x_33); +x_37 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__25; lean_inc(x_1); -x_35 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_1); -x_36 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__24; -x_37 = l_Lean_Expr_const___override(x_36, x_35); -x_38 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__29; -x_39 = l_Lean_Expr_const___override(x_38, x_1); -x_40 = lean_int_neg(x_28); -x_41 = l_Int_toNat(x_40); -lean_dec(x_40); -x_42 = l_Lean_instToExprInt_mkNat(x_41); -x_43 = l_Lean_mkApp3(x_37, x_5, x_39, x_42); -x_44 = l_Lean_Expr_app___override(x_24, x_43); -x_45 = l_Lean_Expr_app___override(x_44, x_6); -x_46 = l_Lean_mkApp4(x_22, x_31, x_45, x_19, x_9); +x_38 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_1); +x_39 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__24; +x_40 = l_Lean_Expr_const___override(x_39, x_38); +x_41 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__29; +x_42 = l_Lean_Expr_const___override(x_41, x_1); +x_43 = lean_int_neg(x_30); +x_44 = l_Int_toNat(x_43); +lean_dec(x_43); +x_45 = l_Lean_instToExprInt_mkNat(x_44); +x_46 = l_Lean_mkApp3(x_40, x_5, x_42, x_45); +x_47 = l_Lean_Expr_app___override(x_26, x_46); +x_48 = l_Lean_Expr_app___override(x_47, x_6); +x_49 = l_Lean_mkApp4(x_24, x_34, x_48, x_21, x_9); +x_50 = lean_box(x_13); +lean_inc(x_18); +lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -x_47 = lean_apply_8(x_7, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_33); -if (lean_obj_tag(x_47) == 0) +x_51 = lean_apply_10(x_7, x_10, x_11, x_12, x_50, x_14, x_15, x_16, x_17, x_18, x_36); +if (lean_obj_tag(x_51) == 0) { -lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; -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_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__3___closed__5; -x_51 = lean_array_push(x_50, x_32); -x_52 = lean_array_push(x_51, x_48); -x_53 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__3___closed__4; +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_52 = lean_ctor_get(x_51, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_51, 1); +lean_inc(x_53); +lean_dec(x_51); +x_54 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__3___closed__5; +x_55 = lean_array_push(x_54, x_35); +x_56 = lean_array_push(x_55, x_52); +x_57 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__3___closed__4; +lean_inc(x_18); +lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -x_54 = l_Lean_Meta_mkAppM(x_53, x_52, x_13, x_14, x_15, x_16, x_49); -if (lean_obj_tag(x_54) == 0) +x_58 = l_Lean_Meta_mkAppM(x_57, x_56, x_15, x_16, x_17, x_18, x_53); +if (lean_obj_tag(x_58) == 0) { -lean_object* x_55; lean_object* x_56; lean_object* x_57; -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); +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); +lean_inc(x_18); +lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -x_57 = l_Lean_Meta_mkEqSymm(x_46, x_13, x_14, x_15, x_16, x_56); -if (lean_obj_tag(x_57) == 0) -{ -lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_58 = lean_ctor_get(x_57, 0); -lean_inc(x_58); -x_59 = lean_ctor_get(x_57, 1); -lean_inc(x_59); -lean_dec(x_57); -x_60 = l_Lean_Meta_mkEqTrans(x_55, x_58, x_13, x_14, x_15, x_16, x_59); -return x_60; -} -else -{ -uint8_t x_61; -lean_dec(x_55); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -x_61 = !lean_is_exclusive(x_57); -if (x_61 == 0) -{ -return x_57; -} -else +x_61 = l_Lean_Meta_mkEqSymm(x_49, x_15, x_16, x_17, x_18, x_60); +if (lean_obj_tag(x_61) == 0) { lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_62 = lean_ctor_get(x_57, 0); -x_63 = lean_ctor_get(x_57, 1); -lean_inc(x_63); +x_62 = lean_ctor_get(x_61, 0); lean_inc(x_62); -lean_dec(x_57); -x_64 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_64, 0, x_62); -lean_ctor_set(x_64, 1, x_63); +x_63 = lean_ctor_get(x_61, 1); +lean_inc(x_63); +lean_dec(x_61); +x_64 = l_Lean_Meta_mkEqTrans(x_59, x_62, x_15, x_16, x_17, x_18, x_63); return x_64; } -} -} else { uint8_t x_65; -lean_dec(x_46); +lean_dec(x_59); +lean_dec(x_18); +lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -x_65 = !lean_is_exclusive(x_54); +x_65 = !lean_is_exclusive(x_61); if (x_65 == 0) { -return x_54; +return x_61; } else { lean_object* x_66; lean_object* x_67; lean_object* x_68; -x_66 = lean_ctor_get(x_54, 0); -x_67 = lean_ctor_get(x_54, 1); +x_66 = lean_ctor_get(x_61, 0); +x_67 = lean_ctor_get(x_61, 1); lean_inc(x_67); lean_inc(x_66); -lean_dec(x_54); +lean_dec(x_61); x_68 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_68, 0, x_66); lean_ctor_set(x_68, 1, x_67); @@ -6502,25 +6553,24 @@ return x_68; else { uint8_t x_69; -lean_dec(x_46); -lean_dec(x_32); +lean_dec(x_49); +lean_dec(x_18); +lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -x_69 = !lean_is_exclusive(x_47); +x_69 = !lean_is_exclusive(x_58); if (x_69 == 0) { -return x_47; +return x_58; } else { lean_object* x_70; lean_object* x_71; lean_object* x_72; -x_70 = lean_ctor_get(x_47, 0); -x_71 = lean_ctor_get(x_47, 1); +x_70 = lean_ctor_get(x_58, 0); +x_71 = lean_ctor_get(x_58, 1); lean_inc(x_71); lean_inc(x_70); -lean_dec(x_47); +lean_dec(x_58); x_72 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_72, 0, x_70); lean_ctor_set(x_72, 1, x_71); @@ -6531,14 +6581,44 @@ return x_72; else { uint8_t x_73; -lean_dec(x_31); +lean_dec(x_49); +lean_dec(x_35); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +x_73 = !lean_is_exclusive(x_51); +if (x_73 == 0) +{ +return x_51; +} +else +{ +lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_74 = lean_ctor_get(x_51, 0); +x_75 = lean_ctor_get(x_51, 1); +lean_inc(x_75); +lean_inc(x_74); +lean_dec(x_51); +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 +{ +uint8_t x_77; +lean_dec(x_34); +lean_dec(x_26); lean_dec(x_24); -lean_dec(x_22); -lean_dec(x_19); +lean_dec(x_21); +lean_dec(x_18); +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); @@ -6547,217 +6627,220 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_1); -x_73 = !lean_is_exclusive(x_30); -if (x_73 == 0) +x_77 = !lean_is_exclusive(x_33); +if (x_77 == 0) { -return x_30; +return x_33; } else { -lean_object* x_74; lean_object* x_75; lean_object* x_76; -x_74 = lean_ctor_get(x_30, 0); -x_75 = lean_ctor_get(x_30, 1); -lean_inc(x_75); -lean_inc(x_74); -lean_dec(x_30); -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_object* x_78; lean_object* x_79; lean_object* x_80; +x_78 = lean_ctor_get(x_33, 0); +x_79 = lean_ctor_get(x_33, 1); +lean_inc(x_79); +lean_inc(x_78); +lean_dec(x_33); +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; } } } -block_116: +block_121: { -if (lean_obj_tag(x_30) == 0) +if (lean_obj_tag(x_33) == 0) { -lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; -x_79 = lean_ctor_get(x_30, 0); -lean_inc(x_79); -x_80 = lean_ctor_get(x_30, 1); -lean_inc(x_80); -lean_dec(x_30); -x_81 = l_Int_toNat(x_28); -x_82 = l_Lean_instToExprInt_mkNat(x_81); -x_83 = l_Lean_Expr_app___override(x_24, x_82); -x_84 = l_Lean_Expr_app___override(x_83, x_6); -x_85 = l_Lean_mkApp4(x_22, x_78, x_84, x_19, x_9); +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_83 = lean_ctor_get(x_33, 0); +lean_inc(x_83); +x_84 = lean_ctor_get(x_33, 1); +lean_inc(x_84); +lean_dec(x_33); +x_85 = l_Int_toNat(x_30); +x_86 = l_Lean_instToExprInt_mkNat(x_85); +x_87 = l_Lean_Expr_app___override(x_26, x_86); +x_88 = l_Lean_Expr_app___override(x_87, x_6); +x_89 = l_Lean_mkApp4(x_24, x_82, x_88, x_21, x_9); +x_90 = lean_box(x_13); +lean_inc(x_18); +lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -x_86 = lean_apply_8(x_7, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_80); -if (lean_obj_tag(x_86) == 0) +x_91 = lean_apply_10(x_7, x_10, x_11, x_12, x_90, x_14, x_15, x_16, x_17, x_18, x_84); +if (lean_obj_tag(x_91) == 0) { -lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; -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_Elab_Tactic_Omega_asLinearComboImpl___lambda__3___closed__5; -x_90 = lean_array_push(x_89, x_79); -x_91 = lean_array_push(x_90, x_87); -x_92 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__3___closed__4; +lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_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_Elab_Tactic_Omega_asLinearComboImpl___lambda__3___closed__5; +x_95 = lean_array_push(x_94, x_83); +x_96 = lean_array_push(x_95, x_92); +x_97 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__3___closed__4; +lean_inc(x_18); +lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -x_93 = l_Lean_Meta_mkAppM(x_92, x_91, x_13, x_14, x_15, x_16, x_88); -if (lean_obj_tag(x_93) == 0) +x_98 = l_Lean_Meta_mkAppM(x_97, x_96, x_15, x_16, x_17, x_18, x_93); +if (lean_obj_tag(x_98) == 0) { -lean_object* x_94; lean_object* x_95; lean_object* x_96; -x_94 = lean_ctor_get(x_93, 0); -lean_inc(x_94); -x_95 = lean_ctor_get(x_93, 1); -lean_inc(x_95); -lean_dec(x_93); +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); +lean_inc(x_18); +lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -x_96 = l_Lean_Meta_mkEqSymm(x_85, x_13, x_14, x_15, x_16, x_95); -if (lean_obj_tag(x_96) == 0) +x_101 = l_Lean_Meta_mkEqSymm(x_89, x_15, x_16, x_17, x_18, x_100); +if (lean_obj_tag(x_101) == 0) { -lean_object* x_97; lean_object* x_98; lean_object* x_99; -x_97 = lean_ctor_get(x_96, 0); -lean_inc(x_97); -x_98 = lean_ctor_get(x_96, 1); -lean_inc(x_98); -lean_dec(x_96); -x_99 = l_Lean_Meta_mkEqTrans(x_94, x_97, x_13, x_14, x_15, x_16, x_98); -return x_99; +lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_102 = lean_ctor_get(x_101, 0); +lean_inc(x_102); +x_103 = lean_ctor_get(x_101, 1); +lean_inc(x_103); +lean_dec(x_101); +x_104 = l_Lean_Meta_mkEqTrans(x_99, x_102, x_15, x_16, x_17, x_18, x_103); +return x_104; } else { -uint8_t x_100; -lean_dec(x_94); +uint8_t x_105; +lean_dec(x_99); +lean_dec(x_18); +lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -x_100 = !lean_is_exclusive(x_96); -if (x_100 == 0) +x_105 = !lean_is_exclusive(x_101); +if (x_105 == 0) { -return x_96; +return x_101; } else { -lean_object* x_101; lean_object* x_102; lean_object* x_103; -x_101 = lean_ctor_get(x_96, 0); -x_102 = lean_ctor_get(x_96, 1); -lean_inc(x_102); -lean_inc(x_101); -lean_dec(x_96); -x_103 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_103, 0, x_101); -lean_ctor_set(x_103, 1, x_102); -return x_103; +lean_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 { -uint8_t x_104; -lean_dec(x_85); +uint8_t x_109; +lean_dec(x_89); +lean_dec(x_18); +lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -x_104 = !lean_is_exclusive(x_93); -if (x_104 == 0) +x_109 = !lean_is_exclusive(x_98); +if (x_109 == 0) { -return x_93; +return x_98; } else { -lean_object* x_105; lean_object* x_106; lean_object* x_107; -x_105 = lean_ctor_get(x_93, 0); -x_106 = lean_ctor_get(x_93, 1); -lean_inc(x_106); -lean_inc(x_105); -lean_dec(x_93); -x_107 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_107, 0, x_105); -lean_ctor_set(x_107, 1, x_106); -return x_107; -} +lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_110 = lean_ctor_get(x_98, 0); +x_111 = lean_ctor_get(x_98, 1); +lean_inc(x_111); +lean_inc(x_110); +lean_dec(x_98); +x_112 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_112, 0, x_110); +lean_ctor_set(x_112, 1, x_111); +return x_112; +} } } else { -uint8_t x_108; -lean_dec(x_85); -lean_dec(x_79); +uint8_t x_113; +lean_dec(x_89); +lean_dec(x_83); +lean_dec(x_18); +lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -x_108 = !lean_is_exclusive(x_86); -if (x_108 == 0) +x_113 = !lean_is_exclusive(x_91); +if (x_113 == 0) { -return x_86; +return x_91; } else { -lean_object* x_109; lean_object* x_110; lean_object* x_111; -x_109 = lean_ctor_get(x_86, 0); -x_110 = lean_ctor_get(x_86, 1); -lean_inc(x_110); -lean_inc(x_109); -lean_dec(x_86); -x_111 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_111, 0, x_109); -lean_ctor_set(x_111, 1, x_110); -return x_111; +lean_object* x_114; lean_object* x_115; lean_object* x_116; +x_114 = lean_ctor_get(x_91, 0); +x_115 = lean_ctor_get(x_91, 1); +lean_inc(x_115); +lean_inc(x_114); +lean_dec(x_91); +x_116 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_116, 0, x_114); +lean_ctor_set(x_116, 1, x_115); +return x_116; } } } else { -uint8_t x_112; -lean_dec(x_78); +uint8_t x_117; +lean_dec(x_82); +lean_dec(x_26); lean_dec(x_24); -lean_dec(x_22); -lean_dec(x_19); +lean_dec(x_21); +lean_dec(x_18); +lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); -lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); -x_112 = !lean_is_exclusive(x_30); -if (x_112 == 0) +x_117 = !lean_is_exclusive(x_33); +if (x_117 == 0) { -return x_30; +return x_33; } else { -lean_object* x_113; lean_object* x_114; lean_object* x_115; -x_113 = lean_ctor_get(x_30, 0); -x_114 = lean_ctor_get(x_30, 1); -lean_inc(x_114); -lean_inc(x_113); -lean_dec(x_30); -x_115 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_115, 0, x_113); -lean_ctor_set(x_115, 1, x_114); -return x_115; +lean_object* x_118; lean_object* x_119; lean_object* x_120; +x_118 = lean_ctor_get(x_33, 0); +x_119 = lean_ctor_get(x_33, 1); +lean_inc(x_119); +lean_inc(x_118); +lean_dec(x_33); +x_120 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_120, 0, x_118); +lean_ctor_set(x_120, 1, x_119); +return x_120; } } } } else { -uint8_t x_133; +uint8_t x_138; +lean_dec(x_18); +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); @@ -6768,23 +6851,23 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_1); -x_133 = !lean_is_exclusive(x_18); -if (x_133 == 0) +x_138 = !lean_is_exclusive(x_20); +if (x_138 == 0) { -return x_18; +return x_20; } else { -lean_object* x_134; lean_object* x_135; lean_object* x_136; -x_134 = lean_ctor_get(x_18, 0); -x_135 = lean_ctor_get(x_18, 1); -lean_inc(x_135); -lean_inc(x_134); -lean_dec(x_18); -x_136 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_136, 0, x_134); -lean_ctor_set(x_136, 1, x_135); -return x_136; +lean_object* x_139; lean_object* x_140; lean_object* x_141; +x_139 = lean_ctor_get(x_20, 0); +x_140 = lean_ctor_get(x_20, 1); +lean_inc(x_140); +lean_inc(x_139); +lean_dec(x_20); +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; } } } @@ -6868,312 +6951,312 @@ lean_ctor_set(x_4, 1, x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___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_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__4(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_object* x_11, lean_object* x_12) { _start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_11 = lean_ctor_get(x_2, 1); -lean_inc(x_11); -x_12 = lean_ctor_get(x_2, 0); -lean_inc(x_12); -lean_dec(x_2); -x_13 = lean_ctor_get(x_11, 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_2, 1); lean_inc(x_13); -x_14 = lean_ctor_get(x_11, 1); +x_14 = lean_ctor_get(x_2, 0); lean_inc(x_14); -if (lean_is_exclusive(x_11)) { - lean_ctor_release(x_11, 0); - lean_ctor_release(x_11, 1); - x_15 = x_11; +lean_dec(x_2); +x_15 = lean_ctor_get(x_13, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +if (lean_is_exclusive(x_13)) { + lean_ctor_release(x_13, 0); + lean_ctor_release(x_13, 1); + x_17 = x_13; } else { - lean_dec_ref(x_11); - x_15 = lean_box(0); + lean_dec_ref(x_13); + x_17 = lean_box(0); } -x_16 = l_Lean_Elab_Tactic_Omega_asLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_16) == 0) +x_18 = l_Lean_Elab_Tactic_Omega_asLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_18) == 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_76; uint8_t x_77; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_17, 1); -lean_inc(x_18); -x_19 = lean_ctor_get(x_16, 1); +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_78; uint8_t x_79; +x_19 = lean_ctor_get(x_18, 0); lean_inc(x_19); -if (lean_is_exclusive(x_16)) { - lean_ctor_release(x_16, 0); - lean_ctor_release(x_16, 1); - x_20 = x_16; -} else { - lean_dec_ref(x_16); - x_20 = lean_box(0); -} -x_21 = lean_ctor_get(x_17, 0); +x_20 = lean_ctor_get(x_19, 1); +lean_inc(x_20); +x_21 = lean_ctor_get(x_18, 1); lean_inc(x_21); -if (lean_is_exclusive(x_17)) { - lean_ctor_release(x_17, 0); - lean_ctor_release(x_17, 1); - x_22 = x_17; -} else { - lean_dec_ref(x_17); - x_22 = lean_box(0); -} -x_23 = lean_ctor_get(x_18, 0); -lean_inc(x_23); -x_24 = lean_ctor_get(x_18, 1); -lean_inc(x_24); if (lean_is_exclusive(x_18)) { lean_ctor_release(x_18, 0); lean_ctor_release(x_18, 1); - x_25 = x_18; + x_22 = x_18; } else { lean_dec_ref(x_18); - x_25 = lean_box(0); + x_22 = lean_box(0); } -x_26 = lean_ctor_get(x_12, 1); -lean_inc(x_26); -x_76 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__4___closed__7; +x_23 = lean_ctor_get(x_19, 0); +lean_inc(x_23); +if (lean_is_exclusive(x_19)) { + lean_ctor_release(x_19, 0); + lean_ctor_release(x_19, 1); + x_24 = x_19; +} else { + lean_dec_ref(x_19); + x_24 = lean_box(0); +} +x_25 = lean_ctor_get(x_20, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_20, 1); lean_inc(x_26); -x_77 = l_List_decidableBAll___rarg(x_76, x_26); -if (x_77 == 0) -{ -lean_object* x_78; uint8_t x_79; -x_78 = lean_ctor_get(x_21, 1); -lean_inc(x_78); -x_79 = l_List_decidableBAll___rarg(x_76, x_78); +if (lean_is_exclusive(x_20)) { + lean_ctor_release(x_20, 0); + lean_ctor_release(x_20, 1); + x_27 = x_20; +} else { + lean_dec_ref(x_20); + x_27 = lean_box(0); +} +x_28 = lean_ctor_get(x_14, 1); +lean_inc(x_28); +x_78 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__4___closed__7; +lean_inc(x_28); +x_79 = l_List_decidableBAll___rarg(x_78, x_28); if (x_79 == 0) { -lean_object* x_80; lean_object* x_81; +lean_object* x_80; uint8_t x_81; +x_80 = lean_ctor_get(x_23, 1); +lean_inc(x_80); +x_81 = l_List_decidableBAll___rarg(x_78, x_80); +if (x_81 == 0) +{ +lean_object* x_82; lean_object* x_83; +lean_dec(x_28); +lean_dec(x_27); lean_dec(x_26); lean_dec(x_25); lean_dec(x_24); lean_dec(x_23); lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_20); +lean_dec(x_17); +lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -x_80 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__4___closed__8; -x_81 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_81, 0, x_80); -lean_ctor_set(x_81, 1, x_19); -return x_81; +x_82 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__4___closed__8; +x_83 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_83, 0, x_82); +lean_ctor_set(x_83, 1, x_21); +return x_83; } else { -lean_object* x_82; -x_82 = lean_box(0); -x_27 = x_82; -goto block_75; +lean_object* x_84; +x_84 = lean_box(0); +x_29 = x_84; +goto block_77; } } else { -lean_object* x_83; -x_83 = lean_box(0); -x_27 = x_83; -goto block_75; +lean_object* x_85; +x_85 = lean_box(0); +x_29 = x_85; +goto block_77; } -block_75: +block_77: { -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; 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_dec(x_27); -x_28 = lean_box(0); -x_29 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__17; -x_30 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__21; -x_31 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1(x_29, x_30, x_26); -lean_dec(x_26); -x_32 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__4___closed__6; -lean_inc(x_31); -x_33 = l_Lean_Expr_app___override(x_32, x_31); -x_34 = lean_ctor_get(x_21, 1); -lean_inc(x_34); -x_35 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1(x_29, x_30, x_34); -lean_dec(x_34); -lean_inc(x_35); -x_36 = l_Lean_Expr_app___override(x_32, x_35); -x_37 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__4___closed__3; -x_38 = l_Lean_mkAppB(x_37, x_33, x_36); -x_39 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__2___boxed), 7, 1); -lean_closure_set(x_39, 0, x_38); -x_40 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); -lean_closure_set(x_40, 0, x_39); -x_41 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; 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; +lean_dec(x_29); +x_30 = lean_box(0); +x_31 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__17; +x_32 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__21; +x_33 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1(x_31, x_32, x_28); +lean_dec(x_28); +x_34 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__4___closed__6; +lean_inc(x_33); +x_35 = l_Lean_Expr_app___override(x_34, x_33); +x_36 = lean_ctor_get(x_23, 1); +lean_inc(x_36); +x_37 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1(x_31, x_32, x_36); +lean_dec(x_36); +lean_inc(x_37); +x_38 = l_Lean_Expr_app___override(x_34, x_37); +x_39 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__4___closed__3; +x_40 = l_Lean_mkAppB(x_39, x_35, x_38); +x_41 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__2___boxed), 9, 1); lean_closure_set(x_41, 0, x_40); -x_42 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__11; -lean_inc(x_21); -lean_inc(x_12); -x_43 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__3___boxed), 17, 8); -lean_closure_set(x_43, 0, x_28); -lean_closure_set(x_43, 1, x_12); -lean_closure_set(x_43, 2, x_21); -lean_closure_set(x_43, 3, x_13); -lean_closure_set(x_43, 4, x_42); -lean_closure_set(x_43, 5, x_35); -lean_closure_set(x_43, 6, x_23); -lean_closure_set(x_43, 7, x_31); -x_44 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg), 10, 2); -lean_closure_set(x_44, 0, x_41); -lean_closure_set(x_44, 1, x_43); -x_45 = l_Lean_Omega_LinearCombo_mul(x_12, x_21); -x_46 = lean_ctor_get(x_24, 1); -lean_inc(x_46); -lean_dec(x_24); -x_47 = lean_array_get_size(x_46); -x_48 = lean_unsigned_to_nat(0u); -x_49 = lean_nat_dec_lt(x_48, x_47); -if (x_49 == 0) +x_42 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); +lean_closure_set(x_42, 0, x_41); +x_43 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); +lean_closure_set(x_43, 0, x_42); +x_44 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__11; +lean_inc(x_23); +lean_inc(x_14); +x_45 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__3___boxed), 19, 8); +lean_closure_set(x_45, 0, x_30); +lean_closure_set(x_45, 1, x_14); +lean_closure_set(x_45, 2, x_23); +lean_closure_set(x_45, 3, x_15); +lean_closure_set(x_45, 4, x_44); +lean_closure_set(x_45, 5, x_37); +lean_closure_set(x_45, 6, x_25); +lean_closure_set(x_45, 7, x_33); +x_46 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg___boxed), 12, 2); +lean_closure_set(x_46, 0, x_43); +lean_closure_set(x_46, 1, x_45); +x_47 = l_Lean_Omega_LinearCombo_mul(x_14, x_23); +x_48 = lean_ctor_get(x_26, 1); +lean_inc(x_48); +lean_dec(x_26); +x_49 = lean_array_get_size(x_48); +x_50 = lean_unsigned_to_nat(0u); +x_51 = lean_nat_dec_lt(x_50, x_49); +if (x_51 == 0) { -lean_object* x_50; lean_object* x_51; lean_object* x_52; uint8_t x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; -lean_dec(x_47); -lean_dec(x_46); -if (lean_is_scalar(x_25)) { - x_50 = lean_alloc_ctor(0, 2, 0); +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; +lean_dec(x_49); +lean_dec(x_48); +if (lean_is_scalar(x_27)) { + x_52 = lean_alloc_ctor(0, 2, 0); } else { - x_50 = x_25; + x_52 = x_27; } -lean_ctor_set(x_50, 0, x_44); -lean_ctor_set(x_50, 1, x_14); -if (lean_is_scalar(x_22)) { - x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_52, 0, x_46); +lean_ctor_set(x_52, 1, x_16); +if (lean_is_scalar(x_24)) { + x_53 = lean_alloc_ctor(0, 2, 0); } else { - x_51 = x_22; + x_53 = x_24; } -lean_ctor_set(x_51, 0, x_45); -lean_ctor_set(x_51, 1, x_50); -x_52 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_52, 0, x_51); -x_53 = 1; -x_54 = lean_box(x_53); -if (lean_is_scalar(x_15)) { - x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_53, 0, x_47); +lean_ctor_set(x_53, 1, x_52); +x_54 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_54, 0, x_53); +x_55 = 1; +x_56 = lean_box(x_55); +if (lean_is_scalar(x_17)) { + x_57 = lean_alloc_ctor(0, 2, 0); } else { - x_55 = x_15; + x_57 = x_17; } -lean_ctor_set(x_55, 0, x_52); -lean_ctor_set(x_55, 1, x_54); -if (lean_is_scalar(x_20)) { - x_56 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_57, 0, x_54); +lean_ctor_set(x_57, 1, x_56); +if (lean_is_scalar(x_22)) { + x_58 = lean_alloc_ctor(0, 2, 0); } else { - x_56 = x_20; + x_58 = x_22; } -lean_ctor_set(x_56, 0, x_55); -lean_ctor_set(x_56, 1, x_19); -return x_56; +lean_ctor_set(x_58, 0, x_57); +lean_ctor_set(x_58, 1, x_21); +return x_58; } else { -uint8_t x_57; -x_57 = lean_nat_dec_le(x_47, x_47); -if (x_57 == 0) +uint8_t x_59; +x_59 = lean_nat_dec_le(x_49, x_49); +if (x_59 == 0) { -lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; -lean_dec(x_47); -lean_dec(x_46); -if (lean_is_scalar(x_25)) { - x_58 = lean_alloc_ctor(0, 2, 0); +lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +lean_dec(x_49); +lean_dec(x_48); +if (lean_is_scalar(x_27)) { + x_60 = lean_alloc_ctor(0, 2, 0); } else { - x_58 = x_25; + x_60 = x_27; } -lean_ctor_set(x_58, 0, x_44); -lean_ctor_set(x_58, 1, x_14); -if (lean_is_scalar(x_22)) { - x_59 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_60, 0, x_46); +lean_ctor_set(x_60, 1, x_16); +if (lean_is_scalar(x_24)) { + x_61 = lean_alloc_ctor(0, 2, 0); } else { - x_59 = x_22; + x_61 = x_24; } -lean_ctor_set(x_59, 0, x_45); -lean_ctor_set(x_59, 1, x_58); -x_60 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_60, 0, x_59); -x_61 = 1; -x_62 = lean_box(x_61); -if (lean_is_scalar(x_15)) { - x_63 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_61, 0, x_47); +lean_ctor_set(x_61, 1, x_60); +x_62 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_62, 0, x_61); +x_63 = 1; +x_64 = lean_box(x_63); +if (lean_is_scalar(x_17)) { + x_65 = lean_alloc_ctor(0, 2, 0); } else { - x_63 = x_15; + x_65 = x_17; } -lean_ctor_set(x_63, 0, x_60); -lean_ctor_set(x_63, 1, x_62); -if (lean_is_scalar(x_20)) { - x_64 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_65, 0, x_62); +lean_ctor_set(x_65, 1, x_64); +if (lean_is_scalar(x_22)) { + x_66 = lean_alloc_ctor(0, 2, 0); } else { - x_64 = x_20; + x_66 = x_22; } -lean_ctor_set(x_64, 0, x_63); -lean_ctor_set(x_64, 1, x_19); -return x_64; +lean_ctor_set(x_66, 0, x_65); +lean_ctor_set(x_66, 1, x_21); +return x_66; } else { -size_t x_65; size_t x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; uint8_t x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; -x_65 = 0; -x_66 = lean_usize_of_nat(x_47); -lean_dec(x_47); -x_67 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Omega_asLinearComboImpl___spec__3(x_46, x_65, x_66, x_14); -lean_dec(x_46); -if (lean_is_scalar(x_25)) { - x_68 = lean_alloc_ctor(0, 2, 0); +size_t x_67; size_t x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; uint8_t x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_67 = 0; +x_68 = lean_usize_of_nat(x_49); +lean_dec(x_49); +x_69 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Omega_asLinearComboImpl___spec__3(x_48, x_67, x_68, x_16); +lean_dec(x_48); +if (lean_is_scalar(x_27)) { + x_70 = lean_alloc_ctor(0, 2, 0); } else { - x_68 = x_25; + x_70 = x_27; } -lean_ctor_set(x_68, 0, x_44); -lean_ctor_set(x_68, 1, x_67); -if (lean_is_scalar(x_22)) { - x_69 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_70, 0, x_46); +lean_ctor_set(x_70, 1, x_69); +if (lean_is_scalar(x_24)) { + x_71 = lean_alloc_ctor(0, 2, 0); } else { - x_69 = x_22; + x_71 = x_24; } -lean_ctor_set(x_69, 0, x_45); -lean_ctor_set(x_69, 1, x_68); -x_70 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_70, 0, x_69); -x_71 = 1; -x_72 = lean_box(x_71); -if (lean_is_scalar(x_15)) { - x_73 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_47); +lean_ctor_set(x_71, 1, x_70); +x_72 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_72, 0, x_71); +x_73 = 1; +x_74 = lean_box(x_73); +if (lean_is_scalar(x_17)) { + x_75 = lean_alloc_ctor(0, 2, 0); } else { - x_73 = x_15; + x_75 = x_17; } -lean_ctor_set(x_73, 0, x_70); -lean_ctor_set(x_73, 1, x_72); -if (lean_is_scalar(x_20)) { - x_74 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_75, 0, x_72); +lean_ctor_set(x_75, 1, x_74); +if (lean_is_scalar(x_22)) { + x_76 = lean_alloc_ctor(0, 2, 0); } else { - x_74 = x_20; + x_76 = x_22; } -lean_ctor_set(x_74, 0, x_73); -lean_ctor_set(x_74, 1, x_19); -return x_74; +lean_ctor_set(x_76, 0, x_75); +lean_ctor_set(x_76, 1, x_21); +return x_76; } } } } else { -uint8_t x_84; +uint8_t x_86; +lean_dec(x_17); +lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -x_84 = !lean_is_exclusive(x_16); -if (x_84 == 0) +x_86 = !lean_is_exclusive(x_18); +if (x_86 == 0) { -return x_16; +return x_18; } else { -lean_object* x_85; lean_object* x_86; lean_object* x_87; -x_85 = lean_ctor_get(x_16, 0); -x_86 = lean_ctor_get(x_16, 1); -lean_inc(x_86); -lean_inc(x_85); -lean_dec(x_16); -x_87 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_87, 0, x_85); -lean_ctor_set(x_87, 1, x_86); -return x_87; +lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_87 = lean_ctor_get(x_18, 0); +x_88 = lean_ctor_get(x_18, 1); +lean_inc(x_88); +lean_inc(x_87); +lean_dec(x_18); +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; } } } @@ -7237,66 +7320,54 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___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_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___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, 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) { _start: { -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; -x_12 = lean_ctor_get(x_1, 0); -x_13 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__8; -x_14 = lean_int_dec_le(x_13, x_12); -x_15 = lean_ctor_get(x_1, 1); -x_16 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__17; -x_17 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__21; -x_18 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1(x_16, x_17, x_15); +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; lean_object* x_22; +x_14 = lean_ctor_get(x_1, 0); +x_15 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__8; +x_16 = lean_int_dec_le(x_15, x_14); +x_17 = lean_ctor_get(x_1, 1); +x_18 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__17; +x_19 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__21; +x_20 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1(x_18, x_19, x_17); +x_21 = lean_box(x_7); +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_19 = lean_apply_8(x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (x_14 == 0) +x_22 = lean_apply_10(x_2, x_4, x_5, x_6, x_21, x_8, x_9, x_10, x_11, x_12, x_13); +if (x_16 == 0) { -if (lean_obj_tag(x_19) == 0) +if (lean_obj_tag(x_22) == 0) { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -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_int_neg(x_12); -x_23 = l_Int_toNat(x_22); +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; 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_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_24 = l_Lean_instToExprInt_mkNat(x_23); -x_25 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__27; -x_26 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__11; -x_27 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__30; -x_28 = l_Lean_mkApp3(x_25, x_26, x_27, x_24); -x_29 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__7; -x_30 = l_Lean_Expr_app___override(x_29, x_28); -x_31 = l_Lean_Expr_app___override(x_30, x_18); -x_32 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__5___closed__3; -x_33 = l_Lean_mkAppB(x_32, x_31, x_3); -x_34 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__5___closed__6; -x_35 = lean_array_push(x_34, x_20); -x_36 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__5___closed__5; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_37 = l_Lean_Meta_mkAppM(x_36, x_35, x_7, x_8, x_9, x_10, x_21); -if (lean_obj_tag(x_37) == 0) -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; -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_25 = lean_int_neg(x_14); +x_26 = l_Int_toNat(x_25); +lean_dec(x_25); +x_27 = l_Lean_instToExprInt_mkNat(x_26); +x_28 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__27; +x_29 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__11; +x_30 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__30; +x_31 = l_Lean_mkApp3(x_28, x_29, x_30, x_27); +x_32 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__7; +x_33 = l_Lean_Expr_app___override(x_32, x_31); +x_34 = l_Lean_Expr_app___override(x_33, x_20); +x_35 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__5___closed__3; +x_36 = l_Lean_mkAppB(x_35, x_34, x_3); +x_37 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__5___closed__6; +x_38 = lean_array_push(x_37, x_23); +x_39 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__5___closed__5; +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_40 = l_Lean_Meta_mkEqSymm(x_33, x_7, x_8, x_9, x_10, x_39); +x_40 = l_Lean_Meta_mkAppM(x_39, x_38, x_9, x_10, x_11, x_12, x_24); if (lean_obj_tag(x_40) == 0) { lean_object* x_41; lean_object* x_42; lean_object* x_43; @@ -7305,132 +7376,132 @@ lean_inc(x_41); x_42 = lean_ctor_get(x_40, 1); lean_inc(x_42); lean_dec(x_40); -x_43 = l_Lean_Meta_mkEqTrans(x_38, x_41, x_7, x_8, x_9, x_10, x_42); -return x_43; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +x_43 = l_Lean_Meta_mkEqSymm(x_36, x_9, x_10, x_11, x_12, x_42); +if (lean_obj_tag(x_43) == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_43, 0); +lean_inc(x_44); +x_45 = lean_ctor_get(x_43, 1); +lean_inc(x_45); +lean_dec(x_43); +x_46 = l_Lean_Meta_mkEqTrans(x_41, x_44, x_9, x_10, x_11, x_12, x_45); +return x_46; } else { -uint8_t x_44; -lean_dec(x_38); +uint8_t x_47; +lean_dec(x_41); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -x_44 = !lean_is_exclusive(x_40); -if (x_44 == 0) +x_47 = !lean_is_exclusive(x_43); +if (x_47 == 0) { -return x_40; +return x_43; } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_45 = lean_ctor_get(x_40, 0); -x_46 = lean_ctor_get(x_40, 1); -lean_inc(x_46); -lean_inc(x_45); -lean_dec(x_40); -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_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_43, 0); +x_49 = lean_ctor_get(x_43, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_43); +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_48; -lean_dec(x_33); +uint8_t x_51; +lean_dec(x_36); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -x_48 = !lean_is_exclusive(x_37); -if (x_48 == 0) +x_51 = !lean_is_exclusive(x_40); +if (x_51 == 0) { -return x_37; +return x_40; } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_49 = lean_ctor_get(x_37, 0); -x_50 = lean_ctor_get(x_37, 1); -lean_inc(x_50); -lean_inc(x_49); -lean_dec(x_37); -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_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; } } } else { -uint8_t x_52; -lean_dec(x_18); +uint8_t x_55; +lean_dec(x_20); +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_3); -x_52 = !lean_is_exclusive(x_19); -if (x_52 == 0) +x_55 = !lean_is_exclusive(x_22); +if (x_55 == 0) { -return x_19; +return x_22; } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_53 = lean_ctor_get(x_19, 0); -x_54 = lean_ctor_get(x_19, 1); -lean_inc(x_54); -lean_inc(x_53); -lean_dec(x_19); -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_56; lean_object* x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_22, 0); +x_57 = lean_ctor_get(x_22, 1); +lean_inc(x_57); +lean_inc(x_56); +lean_dec(x_22); +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 { -if (lean_obj_tag(x_19) == 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; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; -x_56 = lean_ctor_get(x_19, 0); -lean_inc(x_56); -x_57 = lean_ctor_get(x_19, 1); -lean_inc(x_57); -lean_dec(x_19); -x_58 = l_Int_toNat(x_12); -x_59 = l_Lean_instToExprInt_mkNat(x_58); -x_60 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__7; -x_61 = l_Lean_Expr_app___override(x_60, x_59); -x_62 = l_Lean_Expr_app___override(x_61, x_18); -x_63 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__5___closed__3; -x_64 = l_Lean_mkAppB(x_63, x_62, x_3); -x_65 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__5___closed__6; -x_66 = lean_array_push(x_65, x_56); -x_67 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__5___closed__5; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_68 = l_Lean_Meta_mkAppM(x_67, x_66, x_7, x_8, x_9, x_10, x_57); -if (lean_obj_tag(x_68) == 0) +if (lean_obj_tag(x_22) == 0) { -lean_object* x_69; lean_object* x_70; lean_object* x_71; -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); +lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_59 = lean_ctor_get(x_22, 0); +lean_inc(x_59); +x_60 = lean_ctor_get(x_22, 1); +lean_inc(x_60); +lean_dec(x_22); +x_61 = l_Int_toNat(x_14); +x_62 = l_Lean_instToExprInt_mkNat(x_61); +x_63 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__7; +x_64 = l_Lean_Expr_app___override(x_63, x_62); +x_65 = l_Lean_Expr_app___override(x_64, x_20); +x_66 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__5___closed__3; +x_67 = l_Lean_mkAppB(x_66, x_65, x_3); +x_68 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__5___closed__6; +x_69 = lean_array_push(x_68, x_59); +x_70 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__5___closed__5; +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_71 = l_Lean_Meta_mkEqSymm(x_64, x_7, x_8, x_9, x_10, x_70); +x_71 = l_Lean_Meta_mkAppM(x_70, x_69, x_9, x_10, x_11, x_12, x_60); if (lean_obj_tag(x_71) == 0) { lean_object* x_72; lean_object* x_73; lean_object* x_74; @@ -7439,91 +7510,104 @@ lean_inc(x_72); x_73 = lean_ctor_get(x_71, 1); lean_inc(x_73); lean_dec(x_71); -x_74 = l_Lean_Meta_mkEqTrans(x_69, x_72, x_7, x_8, x_9, x_10, x_73); -return x_74; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +x_74 = l_Lean_Meta_mkEqSymm(x_67, x_9, x_10, x_11, x_12, x_73); +if (lean_obj_tag(x_74) == 0) +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_75 = lean_ctor_get(x_74, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_74, 1); +lean_inc(x_76); +lean_dec(x_74); +x_77 = l_Lean_Meta_mkEqTrans(x_72, x_75, x_9, x_10, x_11, x_12, x_76); +return x_77; } else { -uint8_t x_75; -lean_dec(x_69); +uint8_t x_78; +lean_dec(x_72); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -x_75 = !lean_is_exclusive(x_71); -if (x_75 == 0) +x_78 = !lean_is_exclusive(x_74); +if (x_78 == 0) { -return x_71; +return x_74; } 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_79; lean_object* x_80; lean_object* x_81; +x_79 = lean_ctor_get(x_74, 0); +x_80 = lean_ctor_get(x_74, 1); +lean_inc(x_80); +lean_inc(x_79); +lean_dec(x_74); +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 { -uint8_t x_79; -lean_dec(x_64); +uint8_t x_82; +lean_dec(x_67); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -x_79 = !lean_is_exclusive(x_68); -if (x_79 == 0) +x_82 = !lean_is_exclusive(x_71); +if (x_82 == 0) { -return x_68; +return x_71; } else { -lean_object* x_80; lean_object* x_81; lean_object* x_82; -x_80 = lean_ctor_get(x_68, 0); -x_81 = lean_ctor_get(x_68, 1); -lean_inc(x_81); -lean_inc(x_80); -lean_dec(x_68); -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_83; lean_object* x_84; lean_object* x_85; +x_83 = lean_ctor_get(x_71, 0); +x_84 = lean_ctor_get(x_71, 1); +lean_inc(x_84); +lean_inc(x_83); +lean_dec(x_71); +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_83; -lean_dec(x_18); +uint8_t x_86; +lean_dec(x_20); +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_3); -x_83 = !lean_is_exclusive(x_19); -if (x_83 == 0) +x_86 = !lean_is_exclusive(x_22); +if (x_86 == 0) { -return x_19; +return x_22; } else { -lean_object* x_84; lean_object* x_85; lean_object* x_86; -x_84 = lean_ctor_get(x_19, 0); -x_85 = lean_ctor_get(x_19, 1); -lean_inc(x_85); -lean_inc(x_84); -lean_dec(x_19); -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_87; lean_object* x_88; lean_object* x_89; +x_87 = lean_ctor_get(x_22, 0); +x_88 = lean_ctor_get(x_22, 1); +lean_inc(x_88); +lean_inc(x_87); +lean_dec(x_22); +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; } } } @@ -7579,191 +7663,166 @@ x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___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_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___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_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { -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; uint8_t x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_72; -x_14 = lean_ctor_get(x_1, 0); -x_15 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__8; -x_16 = lean_int_dec_le(x_15, x_14); -x_17 = lean_ctor_get(x_1, 1); -x_18 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__17; -x_19 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__21; -x_20 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1(x_18, x_19, x_17); -x_21 = lean_ctor_get(x_2, 0); -x_22 = lean_int_dec_le(x_15, x_21); -x_23 = lean_ctor_get(x_2, 1); -x_24 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1(x_18, x_19, x_23); +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; lean_object* x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_76; +x_16 = lean_ctor_get(x_1, 0); +x_17 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__8; +x_18 = lean_int_dec_le(x_17, x_16); +x_19 = lean_ctor_get(x_1, 1); +x_20 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__17; +x_21 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__21; +x_22 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1(x_20, x_21, x_19); +x_23 = lean_ctor_get(x_2, 0); +x_24 = lean_int_dec_le(x_17, x_23); +x_25 = lean_ctor_get(x_2, 1); +x_26 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1(x_20, x_21, x_25); +x_27 = lean_box(x_9); +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); lean_inc(x_7); lean_inc(x_6); -x_25 = lean_apply_8(x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -if (x_16 == 0) +x_28 = lean_apply_10(x_3, x_6, x_7, x_8, x_27, x_10, x_11, x_12, x_13, x_14, x_15); +if (x_18 == 0) { -lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; -x_113 = lean_int_neg(x_14); -x_114 = l_Int_toNat(x_113); -lean_dec(x_113); -x_115 = l_Lean_instToExprInt_mkNat(x_114); -x_116 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__27; -x_117 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__11; -x_118 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__30; -x_119 = l_Lean_mkApp3(x_116, x_117, x_118, x_115); -x_120 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__7; -x_121 = l_Lean_Expr_app___override(x_120, x_119); -x_122 = l_Lean_Expr_app___override(x_121, x_20); -if (x_22 == 0) +lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_118 = lean_int_neg(x_16); +x_119 = l_Int_toNat(x_118); +lean_dec(x_118); +x_120 = l_Lean_instToExprInt_mkNat(x_119); +x_121 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__27; +x_122 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__11; +x_123 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__30; +x_124 = l_Lean_mkApp3(x_121, x_122, x_123, x_120); +x_125 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__7; +x_126 = l_Lean_Expr_app___override(x_125, x_124); +x_127 = l_Lean_Expr_app___override(x_126, x_22); +if (x_24 == 0) { -x_26 = x_122; -goto block_71; +x_29 = x_127; +goto block_75; } else { -x_72 = x_122; -goto block_112; +x_76 = x_127; +goto block_117; } } else { -lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; -x_123 = l_Int_toNat(x_14); -x_124 = l_Lean_instToExprInt_mkNat(x_123); -x_125 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__7; -x_126 = l_Lean_Expr_app___override(x_125, x_124); -x_127 = l_Lean_Expr_app___override(x_126, x_20); -if (x_22 == 0) +lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; +x_128 = l_Int_toNat(x_16); +x_129 = l_Lean_instToExprInt_mkNat(x_128); +x_130 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__7; +x_131 = l_Lean_Expr_app___override(x_130, x_129); +x_132 = l_Lean_Expr_app___override(x_131, x_22); +if (x_24 == 0) { -x_26 = x_127; -goto block_71; +x_29 = x_132; +goto block_75; } else { -x_72 = x_127; -goto block_112; +x_76 = x_132; +goto block_117; } } -block_71: +block_75: { -if (lean_obj_tag(x_25) == 0) +if (lean_obj_tag(x_28) == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_27 = lean_ctor_get(x_25, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_25, 1); -lean_inc(x_28); -lean_dec(x_25); -x_29 = lean_int_neg(x_21); -x_30 = l_Int_toNat(x_29); -lean_dec(x_29); -x_31 = l_Lean_instToExprInt_mkNat(x_30); -x_32 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__27; -x_33 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__11; -x_34 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__30; -x_35 = l_Lean_mkApp3(x_32, x_33, x_34, x_31); -x_36 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__7; -x_37 = l_Lean_Expr_app___override(x_36, x_35); -x_38 = l_Lean_Expr_app___override(x_37, x_24); -x_39 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__6___closed__3; -x_40 = l_Lean_mkApp3(x_39, x_26, x_38, x_5); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -x_41 = lean_apply_8(x_4, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_28); -if (lean_obj_tag(x_41) == 0) -{ -lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; -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 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__3___closed__5; -x_45 = lean_array_push(x_44, x_27); -x_46 = lean_array_push(x_45, x_42); -x_47 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__6___closed__5; +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; 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_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_int_neg(x_23); +x_33 = l_Int_toNat(x_32); +lean_dec(x_32); +x_34 = l_Lean_instToExprInt_mkNat(x_33); +x_35 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__27; +x_36 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__11; +x_37 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__30; +x_38 = l_Lean_mkApp3(x_35, x_36, x_37, x_34); +x_39 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__7; +x_40 = l_Lean_Expr_app___override(x_39, x_38); +x_41 = l_Lean_Expr_app___override(x_40, x_26); +x_42 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__6___closed__3; +x_43 = l_Lean_mkApp3(x_42, x_29, x_41, x_5); +x_44 = lean_box(x_9); +lean_inc(x_14); +lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -x_48 = l_Lean_Meta_mkAppM(x_47, x_46, x_9, x_10, x_11, x_12, x_43); -if (lean_obj_tag(x_48) == 0) +x_45 = lean_apply_10(x_4, x_6, x_7, x_8, x_44, x_10, x_11, x_12, x_13, x_14, x_31); +if (lean_obj_tag(x_45) == 0) { -lean_object* x_49; lean_object* x_50; lean_object* x_51; -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); +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_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 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__3___closed__5; +x_49 = lean_array_push(x_48, x_30); +x_50 = lean_array_push(x_49, x_46); +x_51 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__6___closed__5; +lean_inc(x_14); +lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -x_51 = l_Lean_Meta_mkEqSymm(x_40, x_9, x_10, x_11, x_12, x_50); -if (lean_obj_tag(x_51) == 0) +x_52 = l_Lean_Meta_mkAppM(x_51, x_50, x_11, x_12, x_13, x_14, x_47); +if (lean_obj_tag(x_52) == 0) { -lean_object* x_52; lean_object* x_53; lean_object* x_54; -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; +x_53 = lean_ctor_get(x_52, 0); lean_inc(x_53); -lean_dec(x_51); -x_54 = l_Lean_Meta_mkEqTrans(x_49, x_52, x_9, x_10, x_11, x_12, x_53); -return x_54; -} -else -{ -uint8_t x_55; -lean_dec(x_49); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -x_55 = !lean_is_exclusive(x_51); -if (x_55 == 0) -{ -return x_51; -} -else +x_54 = lean_ctor_get(x_52, 1); +lean_inc(x_54); +lean_dec(x_52); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +x_55 = l_Lean_Meta_mkEqSymm(x_43, x_11, x_12, x_13, x_14, x_54); +if (lean_obj_tag(x_55) == 0) { lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_56 = lean_ctor_get(x_51, 0); -x_57 = lean_ctor_get(x_51, 1); -lean_inc(x_57); +x_56 = lean_ctor_get(x_55, 0); lean_inc(x_56); -lean_dec(x_51); -x_58 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_58, 0, x_56); -lean_ctor_set(x_58, 1, x_57); +x_57 = lean_ctor_get(x_55, 1); +lean_inc(x_57); +lean_dec(x_55); +x_58 = l_Lean_Meta_mkEqTrans(x_53, x_56, x_11, x_12, x_13, x_14, x_57); return x_58; } -} -} else { uint8_t x_59; -lean_dec(x_40); +lean_dec(x_53); +lean_dec(x_14); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -x_59 = !lean_is_exclusive(x_48); +x_59 = !lean_is_exclusive(x_55); if (x_59 == 0) { -return x_48; +return x_55; } else { lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_48, 0); -x_61 = lean_ctor_get(x_48, 1); +x_60 = lean_ctor_get(x_55, 0); +x_61 = lean_ctor_get(x_55, 1); lean_inc(x_61); lean_inc(x_60); -lean_dec(x_48); +lean_dec(x_55); x_62 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_62, 0, x_60); lean_ctor_set(x_62, 1, x_61); @@ -7774,25 +7833,24 @@ return x_62; else { uint8_t x_63; -lean_dec(x_40); -lean_dec(x_27); +lean_dec(x_43); +lean_dec(x_14); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -x_63 = !lean_is_exclusive(x_41); +x_63 = !lean_is_exclusive(x_52); if (x_63 == 0) { -return x_41; +return x_52; } else { lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_64 = lean_ctor_get(x_41, 0); -x_65 = lean_ctor_get(x_41, 1); +x_64 = lean_ctor_get(x_52, 0); +x_65 = lean_ctor_get(x_52, 1); lean_inc(x_65); lean_inc(x_64); -lean_dec(x_41); +lean_dec(x_52); x_66 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_66, 0, x_64); lean_ctor_set(x_66, 1, x_65); @@ -7803,30 +7861,25 @@ return x_66; else { uint8_t x_67; -lean_dec(x_26); -lean_dec(x_24); +lean_dec(x_43); +lean_dec(x_30); +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); -x_67 = !lean_is_exclusive(x_25); +x_67 = !lean_is_exclusive(x_45); if (x_67 == 0) { -return x_25; +return x_45; } else { lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_68 = lean_ctor_get(x_25, 0); -x_69 = lean_ctor_get(x_25, 1); +x_68 = lean_ctor_get(x_45, 0); +x_69 = lean_ctor_get(x_45, 1); lean_inc(x_69); lean_inc(x_68); -lean_dec(x_25); +lean_dec(x_45); x_70 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_70, 0, x_68); lean_ctor_set(x_70, 1, x_69); @@ -7834,185 +7887,222 @@ return x_70; } } } -block_112: +else { -if (lean_obj_tag(x_25) == 0) +uint8_t x_71; +lean_dec(x_29); +lean_dec(x_26); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_71 = !lean_is_exclusive(x_28); +if (x_71 == 0) +{ +return x_28; +} +else { -lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; -x_73 = lean_ctor_get(x_25, 0); +lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_72 = lean_ctor_get(x_28, 0); +x_73 = lean_ctor_get(x_28, 1); lean_inc(x_73); -x_74 = lean_ctor_get(x_25, 1); -lean_inc(x_74); -lean_dec(x_25); -x_75 = l_Int_toNat(x_21); -x_76 = l_Lean_instToExprInt_mkNat(x_75); -x_77 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__7; -x_78 = l_Lean_Expr_app___override(x_77, x_76); -x_79 = l_Lean_Expr_app___override(x_78, x_24); -x_80 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__6___closed__3; -x_81 = l_Lean_mkApp3(x_80, x_72, x_79, x_5); +lean_inc(x_72); +lean_dec(x_28); +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; +} +} +} +block_117: +{ +if (lean_obj_tag(x_28) == 0) +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_77 = lean_ctor_get(x_28, 0); +lean_inc(x_77); +x_78 = lean_ctor_get(x_28, 1); +lean_inc(x_78); +lean_dec(x_28); +x_79 = l_Int_toNat(x_23); +x_80 = l_Lean_instToExprInt_mkNat(x_79); +x_81 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__7; +x_82 = l_Lean_Expr_app___override(x_81, x_80); +x_83 = l_Lean_Expr_app___override(x_82, x_26); +x_84 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__6___closed__3; +x_85 = l_Lean_mkApp3(x_84, x_76, x_83, x_5); +x_86 = lean_box(x_9); +lean_inc(x_14); +lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -x_82 = lean_apply_8(x_4, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_74); -if (lean_obj_tag(x_82) == 0) +x_87 = lean_apply_10(x_4, x_6, x_7, x_8, x_86, x_10, x_11, x_12, x_13, x_14, x_78); +if (lean_obj_tag(x_87) == 0) { -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_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_Elab_Tactic_Omega_asLinearComboImpl___lambda__3___closed__5; -x_86 = lean_array_push(x_85, x_73); -x_87 = lean_array_push(x_86, x_83); -x_88 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__6___closed__5; +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); +x_90 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__3___closed__5; +x_91 = lean_array_push(x_90, x_77); +x_92 = lean_array_push(x_91, x_88); +x_93 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__6___closed__5; +lean_inc(x_14); +lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -x_89 = l_Lean_Meta_mkAppM(x_88, x_87, x_9, x_10, x_11, x_12, x_84); -if (lean_obj_tag(x_89) == 0) +x_94 = l_Lean_Meta_mkAppM(x_93, x_92, x_11, x_12, x_13, x_14, x_89); +if (lean_obj_tag(x_94) == 0) { -lean_object* x_90; lean_object* x_91; lean_object* x_92; -x_90 = lean_ctor_get(x_89, 0); -lean_inc(x_90); -x_91 = lean_ctor_get(x_89, 1); -lean_inc(x_91); -lean_dec(x_89); +lean_object* x_95; lean_object* x_96; lean_object* x_97; +x_95 = lean_ctor_get(x_94, 0); +lean_inc(x_95); +x_96 = lean_ctor_get(x_94, 1); +lean_inc(x_96); +lean_dec(x_94); +lean_inc(x_14); +lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -x_92 = l_Lean_Meta_mkEqSymm(x_81, x_9, x_10, x_11, x_12, x_91); -if (lean_obj_tag(x_92) == 0) +x_97 = l_Lean_Meta_mkEqSymm(x_85, x_11, x_12, x_13, x_14, x_96); +if (lean_obj_tag(x_97) == 0) { -lean_object* x_93; lean_object* x_94; lean_object* x_95; -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 = l_Lean_Meta_mkEqTrans(x_90, x_93, x_9, x_10, x_11, x_12, x_94); -return x_95; +lean_object* x_98; lean_object* x_99; lean_object* x_100; +x_98 = lean_ctor_get(x_97, 0); +lean_inc(x_98); +x_99 = lean_ctor_get(x_97, 1); +lean_inc(x_99); +lean_dec(x_97); +x_100 = l_Lean_Meta_mkEqTrans(x_95, x_98, x_11, x_12, x_13, x_14, x_99); +return x_100; } else { -uint8_t x_96; -lean_dec(x_90); +uint8_t x_101; +lean_dec(x_95); +lean_dec(x_14); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -x_96 = !lean_is_exclusive(x_92); -if (x_96 == 0) +x_101 = !lean_is_exclusive(x_97); +if (x_101 == 0) { -return x_92; +return x_97; } else { -lean_object* x_97; lean_object* x_98; lean_object* x_99; -x_97 = lean_ctor_get(x_92, 0); -x_98 = lean_ctor_get(x_92, 1); -lean_inc(x_98); -lean_inc(x_97); -lean_dec(x_92); -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_object* x_102; lean_object* x_103; lean_object* x_104; +x_102 = lean_ctor_get(x_97, 0); +x_103 = lean_ctor_get(x_97, 1); +lean_inc(x_103); +lean_inc(x_102); +lean_dec(x_97); +x_104 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_104, 0, x_102); +lean_ctor_set(x_104, 1, x_103); +return x_104; } } } else { -uint8_t x_100; -lean_dec(x_81); +uint8_t x_105; +lean_dec(x_85); +lean_dec(x_14); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -x_100 = !lean_is_exclusive(x_89); -if (x_100 == 0) +x_105 = !lean_is_exclusive(x_94); +if (x_105 == 0) { -return x_89; +return x_94; } else { -lean_object* x_101; lean_object* x_102; lean_object* x_103; -x_101 = lean_ctor_get(x_89, 0); -x_102 = lean_ctor_get(x_89, 1); -lean_inc(x_102); -lean_inc(x_101); -lean_dec(x_89); -x_103 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_103, 0, x_101); -lean_ctor_set(x_103, 1, x_102); -return x_103; +lean_object* x_106; lean_object* x_107; lean_object* x_108; +x_106 = lean_ctor_get(x_94, 0); +x_107 = lean_ctor_get(x_94, 1); +lean_inc(x_107); +lean_inc(x_106); +lean_dec(x_94); +x_108 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_108, 0, x_106); +lean_ctor_set(x_108, 1, x_107); +return x_108; } } } else { -uint8_t x_104; -lean_dec(x_81); -lean_dec(x_73); +uint8_t x_109; +lean_dec(x_85); +lean_dec(x_77); +lean_dec(x_14); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -x_104 = !lean_is_exclusive(x_82); -if (x_104 == 0) +x_109 = !lean_is_exclusive(x_87); +if (x_109 == 0) { -return x_82; +return x_87; } else { -lean_object* x_105; lean_object* x_106; lean_object* x_107; -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_107 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_107, 0, x_105); -lean_ctor_set(x_107, 1, x_106); -return x_107; +lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_110 = lean_ctor_get(x_87, 0); +x_111 = lean_ctor_get(x_87, 1); +lean_inc(x_111); +lean_inc(x_110); +lean_dec(x_87); +x_112 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_112, 0, x_110); +lean_ctor_set(x_112, 1, x_111); +return x_112; } } } else { -uint8_t x_108; -lean_dec(x_72); -lean_dec(x_24); +uint8_t x_113; +lean_dec(x_76); +lean_dec(x_26); +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); -x_108 = !lean_is_exclusive(x_25); -if (x_108 == 0) +x_113 = !lean_is_exclusive(x_28); +if (x_113 == 0) { -return x_25; +return x_28; } else { -lean_object* x_109; lean_object* x_110; lean_object* x_111; -x_109 = lean_ctor_get(x_25, 0); -x_110 = lean_ctor_get(x_25, 1); -lean_inc(x_110); -lean_inc(x_109); -lean_dec(x_25); -x_111 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_111, 0, x_109); -lean_ctor_set(x_111, 1, x_110); -return x_111; +lean_object* x_114; lean_object* x_115; lean_object* x_116; +x_114 = lean_ctor_get(x_28, 0); +x_115 = lean_ctor_get(x_28, 1); +lean_inc(x_115); +lean_inc(x_114); +lean_dec(x_28); +x_116 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_116, 0, x_114); +lean_ctor_set(x_116, 1, x_115); +return x_116; } } } @@ -8068,191 +8158,166 @@ x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___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, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___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, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { -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; uint8_t x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_72; -x_14 = lean_ctor_get(x_1, 0); -x_15 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__8; -x_16 = lean_int_dec_le(x_15, x_14); -x_17 = lean_ctor_get(x_1, 1); -x_18 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__17; -x_19 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__21; -x_20 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1(x_18, x_19, x_17); -x_21 = lean_ctor_get(x_2, 0); -x_22 = lean_int_dec_le(x_15, x_21); -x_23 = lean_ctor_get(x_2, 1); -x_24 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1(x_18, x_19, x_23); +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; lean_object* x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_76; +x_16 = lean_ctor_get(x_1, 0); +x_17 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__8; +x_18 = lean_int_dec_le(x_17, x_16); +x_19 = lean_ctor_get(x_1, 1); +x_20 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__17; +x_21 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__21; +x_22 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1(x_20, x_21, x_19); +x_23 = lean_ctor_get(x_2, 0); +x_24 = lean_int_dec_le(x_17, x_23); +x_25 = lean_ctor_get(x_2, 1); +x_26 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Tactic_Omega_instToExprLinearCombo___spec__1(x_20, x_21, x_25); +x_27 = lean_box(x_9); +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); lean_inc(x_7); lean_inc(x_6); -x_25 = lean_apply_8(x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -if (x_16 == 0) -{ -lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; -x_113 = lean_int_neg(x_14); -x_114 = l_Int_toNat(x_113); -lean_dec(x_113); -x_115 = l_Lean_instToExprInt_mkNat(x_114); -x_116 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__27; -x_117 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__11; -x_118 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__30; -x_119 = l_Lean_mkApp3(x_116, x_117, x_118, x_115); -x_120 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__7; -x_121 = l_Lean_Expr_app___override(x_120, x_119); -x_122 = l_Lean_Expr_app___override(x_121, x_20); -if (x_22 == 0) +x_28 = lean_apply_10(x_3, x_6, x_7, x_8, x_27, x_10, x_11, x_12, x_13, x_14, x_15); +if (x_18 == 0) { -x_26 = x_122; -goto block_71; -} +lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_118 = lean_int_neg(x_16); +x_119 = l_Int_toNat(x_118); +lean_dec(x_118); +x_120 = l_Lean_instToExprInt_mkNat(x_119); +x_121 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__27; +x_122 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__11; +x_123 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__30; +x_124 = l_Lean_mkApp3(x_121, x_122, x_123, x_120); +x_125 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__7; +x_126 = l_Lean_Expr_app___override(x_125, x_124); +x_127 = l_Lean_Expr_app___override(x_126, x_22); +if (x_24 == 0) +{ +x_29 = x_127; +goto block_75; +} else { -x_72 = x_122; -goto block_112; +x_76 = x_127; +goto block_117; } } else { -lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; -x_123 = l_Int_toNat(x_14); -x_124 = l_Lean_instToExprInt_mkNat(x_123); -x_125 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__7; -x_126 = l_Lean_Expr_app___override(x_125, x_124); -x_127 = l_Lean_Expr_app___override(x_126, x_20); -if (x_22 == 0) +lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; +x_128 = l_Int_toNat(x_16); +x_129 = l_Lean_instToExprInt_mkNat(x_128); +x_130 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__7; +x_131 = l_Lean_Expr_app___override(x_130, x_129); +x_132 = l_Lean_Expr_app___override(x_131, x_22); +if (x_24 == 0) { -x_26 = x_127; -goto block_71; +x_29 = x_132; +goto block_75; } else { -x_72 = x_127; -goto block_112; +x_76 = x_132; +goto block_117; } } -block_71: +block_75: { -if (lean_obj_tag(x_25) == 0) +if (lean_obj_tag(x_28) == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_27 = lean_ctor_get(x_25, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_25, 1); -lean_inc(x_28); -lean_dec(x_25); -x_29 = lean_int_neg(x_21); -x_30 = l_Int_toNat(x_29); -lean_dec(x_29); -x_31 = l_Lean_instToExprInt_mkNat(x_30); -x_32 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__27; -x_33 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__11; -x_34 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__30; -x_35 = l_Lean_mkApp3(x_32, x_33, x_34, x_31); -x_36 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__7; -x_37 = l_Lean_Expr_app___override(x_36, x_35); -x_38 = l_Lean_Expr_app___override(x_37, x_24); -x_39 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__7___closed__3; -x_40 = l_Lean_mkApp3(x_39, x_26, x_38, x_5); +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; 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_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_int_neg(x_23); +x_33 = l_Int_toNat(x_32); +lean_dec(x_32); +x_34 = l_Lean_instToExprInt_mkNat(x_33); +x_35 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__27; +x_36 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__11; +x_37 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__30; +x_38 = l_Lean_mkApp3(x_35, x_36, x_37, x_34); +x_39 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__7; +x_40 = l_Lean_Expr_app___override(x_39, x_38); +x_41 = l_Lean_Expr_app___override(x_40, x_26); +x_42 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__7___closed__3; +x_43 = l_Lean_mkApp3(x_42, x_29, x_41, x_5); +x_44 = lean_box(x_9); +lean_inc(x_14); +lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -x_41 = lean_apply_8(x_4, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_28); -if (lean_obj_tag(x_41) == 0) +x_45 = lean_apply_10(x_4, x_6, x_7, x_8, x_44, x_10, x_11, x_12, x_13, x_14, x_31); +if (lean_obj_tag(x_45) == 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; -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 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__3___closed__5; -x_45 = lean_array_push(x_44, x_27); -x_46 = lean_array_push(x_45, x_42); -x_47 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__7___closed__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; lean_object* x_52; +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 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__3___closed__5; +x_49 = lean_array_push(x_48, x_30); +x_50 = lean_array_push(x_49, x_46); +x_51 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__7___closed__5; +lean_inc(x_14); +lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -x_48 = l_Lean_Meta_mkAppM(x_47, x_46, x_9, x_10, x_11, x_12, x_43); -if (lean_obj_tag(x_48) == 0) +x_52 = l_Lean_Meta_mkAppM(x_51, x_50, x_11, x_12, x_13, x_14, x_47); +if (lean_obj_tag(x_52) == 0) { -lean_object* x_49; lean_object* x_50; lean_object* x_51; -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); +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); +lean_inc(x_14); +lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -x_51 = l_Lean_Meta_mkEqSymm(x_40, x_9, x_10, x_11, x_12, x_50); -if (lean_obj_tag(x_51) == 0) -{ -lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_ctor_get(x_51, 0); -lean_inc(x_52); -x_53 = lean_ctor_get(x_51, 1); -lean_inc(x_53); -lean_dec(x_51); -x_54 = l_Lean_Meta_mkEqTrans(x_49, x_52, x_9, x_10, x_11, x_12, x_53); -return x_54; -} -else -{ -uint8_t x_55; -lean_dec(x_49); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -x_55 = !lean_is_exclusive(x_51); -if (x_55 == 0) -{ -return x_51; -} -else +x_55 = l_Lean_Meta_mkEqSymm(x_43, x_11, x_12, x_13, x_14, x_54); +if (lean_obj_tag(x_55) == 0) { lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_56 = lean_ctor_get(x_51, 0); -x_57 = lean_ctor_get(x_51, 1); -lean_inc(x_57); +x_56 = lean_ctor_get(x_55, 0); lean_inc(x_56); -lean_dec(x_51); -x_58 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_58, 0, x_56); -lean_ctor_set(x_58, 1, x_57); +x_57 = lean_ctor_get(x_55, 1); +lean_inc(x_57); +lean_dec(x_55); +x_58 = l_Lean_Meta_mkEqTrans(x_53, x_56, x_11, x_12, x_13, x_14, x_57); return x_58; } -} -} else { uint8_t x_59; -lean_dec(x_40); +lean_dec(x_53); +lean_dec(x_14); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -x_59 = !lean_is_exclusive(x_48); +x_59 = !lean_is_exclusive(x_55); if (x_59 == 0) { -return x_48; +return x_55; } else { lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_48, 0); -x_61 = lean_ctor_get(x_48, 1); +x_60 = lean_ctor_get(x_55, 0); +x_61 = lean_ctor_get(x_55, 1); lean_inc(x_61); lean_inc(x_60); -lean_dec(x_48); +lean_dec(x_55); x_62 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_62, 0, x_60); lean_ctor_set(x_62, 1, x_61); @@ -8263,25 +8328,24 @@ return x_62; else { uint8_t x_63; -lean_dec(x_40); -lean_dec(x_27); +lean_dec(x_43); +lean_dec(x_14); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -x_63 = !lean_is_exclusive(x_41); +x_63 = !lean_is_exclusive(x_52); if (x_63 == 0) { -return x_41; +return x_52; } else { lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_64 = lean_ctor_get(x_41, 0); -x_65 = lean_ctor_get(x_41, 1); +x_64 = lean_ctor_get(x_52, 0); +x_65 = lean_ctor_get(x_52, 1); lean_inc(x_65); lean_inc(x_64); -lean_dec(x_41); +lean_dec(x_52); x_66 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_66, 0, x_64); lean_ctor_set(x_66, 1, x_65); @@ -8292,30 +8356,25 @@ return x_66; else { uint8_t x_67; -lean_dec(x_26); -lean_dec(x_24); +lean_dec(x_43); +lean_dec(x_30); +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); -x_67 = !lean_is_exclusive(x_25); +x_67 = !lean_is_exclusive(x_45); if (x_67 == 0) { -return x_25; +return x_45; } else { lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_68 = lean_ctor_get(x_25, 0); -x_69 = lean_ctor_get(x_25, 1); +x_68 = lean_ctor_get(x_45, 0); +x_69 = lean_ctor_get(x_45, 1); lean_inc(x_69); lean_inc(x_68); -lean_dec(x_25); +lean_dec(x_45); x_70 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_70, 0, x_68); lean_ctor_set(x_70, 1, x_69); @@ -8323,192 +8382,229 @@ return x_70; } } } -block_112: +else { -if (lean_obj_tag(x_25) == 0) +uint8_t x_71; +lean_dec(x_29); +lean_dec(x_26); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_71 = !lean_is_exclusive(x_28); +if (x_71 == 0) +{ +return x_28; +} +else { -lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; -x_73 = lean_ctor_get(x_25, 0); +lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_72 = lean_ctor_get(x_28, 0); +x_73 = lean_ctor_get(x_28, 1); lean_inc(x_73); -x_74 = lean_ctor_get(x_25, 1); -lean_inc(x_74); -lean_dec(x_25); -x_75 = l_Int_toNat(x_21); -x_76 = l_Lean_instToExprInt_mkNat(x_75); -x_77 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__7; -x_78 = l_Lean_Expr_app___override(x_77, x_76); -x_79 = l_Lean_Expr_app___override(x_78, x_24); -x_80 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__7___closed__3; -x_81 = l_Lean_mkApp3(x_80, x_72, x_79, x_5); +lean_inc(x_72); +lean_dec(x_28); +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; +} +} +} +block_117: +{ +if (lean_obj_tag(x_28) == 0) +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_77 = lean_ctor_get(x_28, 0); +lean_inc(x_77); +x_78 = lean_ctor_get(x_28, 1); +lean_inc(x_78); +lean_dec(x_28); +x_79 = l_Int_toNat(x_23); +x_80 = l_Lean_instToExprInt_mkNat(x_79); +x_81 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__7; +x_82 = l_Lean_Expr_app___override(x_81, x_80); +x_83 = l_Lean_Expr_app___override(x_82, x_26); +x_84 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__7___closed__3; +x_85 = l_Lean_mkApp3(x_84, x_76, x_83, x_5); +x_86 = lean_box(x_9); +lean_inc(x_14); +lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -x_82 = lean_apply_8(x_4, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_74); -if (lean_obj_tag(x_82) == 0) +x_87 = lean_apply_10(x_4, x_6, x_7, x_8, x_86, x_10, x_11, x_12, x_13, x_14, x_78); +if (lean_obj_tag(x_87) == 0) { -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_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_Elab_Tactic_Omega_asLinearComboImpl___lambda__3___closed__5; -x_86 = lean_array_push(x_85, x_73); -x_87 = lean_array_push(x_86, x_83); -x_88 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__7___closed__5; +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); +x_90 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__3___closed__5; +x_91 = lean_array_push(x_90, x_77); +x_92 = lean_array_push(x_91, x_88); +x_93 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__7___closed__5; +lean_inc(x_14); +lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -x_89 = l_Lean_Meta_mkAppM(x_88, x_87, x_9, x_10, x_11, x_12, x_84); -if (lean_obj_tag(x_89) == 0) +x_94 = l_Lean_Meta_mkAppM(x_93, x_92, x_11, x_12, x_13, x_14, x_89); +if (lean_obj_tag(x_94) == 0) { -lean_object* x_90; lean_object* x_91; lean_object* x_92; -x_90 = lean_ctor_get(x_89, 0); -lean_inc(x_90); -x_91 = lean_ctor_get(x_89, 1); -lean_inc(x_91); -lean_dec(x_89); +lean_object* x_95; lean_object* x_96; lean_object* x_97; +x_95 = lean_ctor_get(x_94, 0); +lean_inc(x_95); +x_96 = lean_ctor_get(x_94, 1); +lean_inc(x_96); +lean_dec(x_94); +lean_inc(x_14); +lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -x_92 = l_Lean_Meta_mkEqSymm(x_81, x_9, x_10, x_11, x_12, x_91); -if (lean_obj_tag(x_92) == 0) +x_97 = l_Lean_Meta_mkEqSymm(x_85, x_11, x_12, x_13, x_14, x_96); +if (lean_obj_tag(x_97) == 0) { -lean_object* x_93; lean_object* x_94; lean_object* x_95; -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 = l_Lean_Meta_mkEqTrans(x_90, x_93, x_9, x_10, x_11, x_12, x_94); -return x_95; +lean_object* x_98; lean_object* x_99; lean_object* x_100; +x_98 = lean_ctor_get(x_97, 0); +lean_inc(x_98); +x_99 = lean_ctor_get(x_97, 1); +lean_inc(x_99); +lean_dec(x_97); +x_100 = l_Lean_Meta_mkEqTrans(x_95, x_98, x_11, x_12, x_13, x_14, x_99); +return x_100; } else { -uint8_t x_96; -lean_dec(x_90); +uint8_t x_101; +lean_dec(x_95); +lean_dec(x_14); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -x_96 = !lean_is_exclusive(x_92); -if (x_96 == 0) +x_101 = !lean_is_exclusive(x_97); +if (x_101 == 0) { -return x_92; +return x_97; } else { -lean_object* x_97; lean_object* x_98; lean_object* x_99; -x_97 = lean_ctor_get(x_92, 0); -x_98 = lean_ctor_get(x_92, 1); -lean_inc(x_98); -lean_inc(x_97); -lean_dec(x_92); -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_object* x_102; lean_object* x_103; lean_object* x_104; +x_102 = lean_ctor_get(x_97, 0); +x_103 = lean_ctor_get(x_97, 1); +lean_inc(x_103); +lean_inc(x_102); +lean_dec(x_97); +x_104 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_104, 0, x_102); +lean_ctor_set(x_104, 1, x_103); +return x_104; } } } else { -uint8_t x_100; -lean_dec(x_81); +uint8_t x_105; +lean_dec(x_85); +lean_dec(x_14); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -x_100 = !lean_is_exclusive(x_89); -if (x_100 == 0) +x_105 = !lean_is_exclusive(x_94); +if (x_105 == 0) { -return x_89; +return x_94; } else { -lean_object* x_101; lean_object* x_102; lean_object* x_103; -x_101 = lean_ctor_get(x_89, 0); -x_102 = lean_ctor_get(x_89, 1); -lean_inc(x_102); -lean_inc(x_101); -lean_dec(x_89); -x_103 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_103, 0, x_101); -lean_ctor_set(x_103, 1, x_102); -return x_103; +lean_object* x_106; lean_object* x_107; lean_object* x_108; +x_106 = lean_ctor_get(x_94, 0); +x_107 = lean_ctor_get(x_94, 1); +lean_inc(x_107); +lean_inc(x_106); +lean_dec(x_94); +x_108 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_108, 0, x_106); +lean_ctor_set(x_108, 1, x_107); +return x_108; } } } else { -uint8_t x_104; -lean_dec(x_81); -lean_dec(x_73); +uint8_t x_109; +lean_dec(x_85); +lean_dec(x_77); +lean_dec(x_14); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -x_104 = !lean_is_exclusive(x_82); -if (x_104 == 0) +x_109 = !lean_is_exclusive(x_87); +if (x_109 == 0) { -return x_82; +return x_87; } else { -lean_object* x_105; lean_object* x_106; lean_object* x_107; -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_107 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_107, 0, x_105); -lean_ctor_set(x_107, 1, x_106); -return x_107; +lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_110 = lean_ctor_get(x_87, 0); +x_111 = lean_ctor_get(x_87, 1); +lean_inc(x_111); +lean_inc(x_110); +lean_dec(x_87); +x_112 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_112, 0, x_110); +lean_ctor_set(x_112, 1, x_111); +return x_112; } } } else { -uint8_t x_108; -lean_dec(x_72); -lean_dec(x_24); +uint8_t x_113; +lean_dec(x_76); +lean_dec(x_26); +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); -x_108 = !lean_is_exclusive(x_25); -if (x_108 == 0) +x_113 = !lean_is_exclusive(x_28); +if (x_113 == 0) { -return x_25; +return x_28; } else { -lean_object* x_109; lean_object* x_110; lean_object* x_111; -x_109 = lean_ctor_get(x_25, 0); -x_110 = lean_ctor_get(x_25, 1); -lean_inc(x_110); -lean_inc(x_109); -lean_dec(x_25); -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; -} -} -} -} -} -static lean_object* _init_l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__1() { -_start: +lean_object* x_114; lean_object* x_115; lean_object* x_116; +x_114 = lean_ctor_get(x_28, 0); +x_115 = lean_ctor_get(x_28, 1); +lean_inc(x_115); +lean_inc(x_114); +lean_dec(x_28); +x_116 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_116, 0, x_114); +lean_ctor_set(x_116, 1, x_115); +return x_116; +} +} +} +} +} +static lean_object* _init_l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__1() { +_start: { lean_object* x_1; x_1 = lean_mk_string_from_bytes("HSub", 4); @@ -8929,355 +9025,355 @@ x_1 = lean_mk_string_from_bytes("hSub", 4); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___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_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8(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_object* x_11, lean_object* x_12) { _start: { -lean_object* x_11; +lean_object* x_13; lean_dec(x_2); lean_inc(x_1); -x_11 = l_Lean_Elab_Tactic_Omega_groundInt_x3f(x_1); -if (lean_obj_tag(x_11) == 0) +x_13 = l_Lean_Elab_Tactic_Omega_groundInt_x3f(x_1); +if (lean_obj_tag(x_13) == 0) { -uint8_t x_12; -x_12 = l_Lean_Expr_isFVar(x_1); -if (x_12 == 0) +uint8_t x_14; +x_14 = l_Lean_Expr_isFVar(x_1); +if (x_14 == 0) { -lean_object* x_13; lean_object* x_14; +lean_object* x_15; lean_object* x_16; lean_inc(x_1); -x_13 = l_Lean_Expr_getAppFnArgs(x_1); -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -if (lean_obj_tag(x_14) == 1) -{ -lean_object* x_15; -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -if (lean_obj_tag(x_15) == 1) -{ -lean_object* x_16; +x_15 = l_Lean_Expr_getAppFnArgs(x_1); x_16 = lean_ctor_get(x_15, 0); lean_inc(x_16); -if (lean_obj_tag(x_16) == 0) +if (lean_obj_tag(x_16) == 1) { -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_13, 1); +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 0); lean_inc(x_17); -lean_dec(x_13); -x_18 = lean_ctor_get(x_14, 1); +if (lean_obj_tag(x_17) == 1) +{ +lean_object* x_18; +x_18 = lean_ctor_get(x_17, 0); lean_inc(x_18); -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; uint8_t x_23; x_19 = lean_ctor_get(x_15, 1); lean_inc(x_19); lean_dec(x_15); -x_20 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__2; -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_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__1; -x_23 = lean_string_dec_eq(x_19, x_22); +x_20 = lean_ctor_get(x_16, 1); +lean_inc(x_20); +lean_dec(x_16); +x_21 = lean_ctor_get(x_17, 1); +lean_inc(x_21); +lean_dec(x_17); +x_22 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__2; +x_23 = lean_string_dec_eq(x_21, x_22); if (x_23 == 0) { lean_object* x_24; uint8_t x_25; -x_24 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__22; -x_25 = lean_string_dec_eq(x_19, x_24); +x_24 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__1; +x_25 = lean_string_dec_eq(x_21, x_24); if (x_25 == 0) { lean_object* x_26; uint8_t x_27; -x_26 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__2; -x_27 = lean_string_dec_eq(x_19, x_26); +x_26 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__22; +x_27 = lean_string_dec_eq(x_21, x_26); if (x_27 == 0) { lean_object* x_28; uint8_t x_29; -x_28 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__3; -x_29 = lean_string_dec_eq(x_19, x_28); +x_28 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__2; +x_29 = lean_string_dec_eq(x_21, x_28); if (x_29 == 0) { lean_object* x_30; uint8_t x_31; -x_30 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__4; -x_31 = lean_string_dec_eq(x_19, x_30); +x_30 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__3; +x_31 = lean_string_dec_eq(x_21, x_30); if (x_31 == 0) { lean_object* x_32; uint8_t x_33; -x_32 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__5; -x_33 = lean_string_dec_eq(x_19, x_32); +x_32 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__4; +x_33 = lean_string_dec_eq(x_21, x_32); if (x_33 == 0) { lean_object* x_34; uint8_t x_35; -x_34 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__6; -x_35 = lean_string_dec_eq(x_19, x_34); +x_34 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__5; +x_35 = lean_string_dec_eq(x_21, x_34); if (x_35 == 0) { lean_object* x_36; uint8_t x_37; -x_36 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__7; -x_37 = lean_string_dec_eq(x_19, x_36); +x_36 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__6; +x_37 = lean_string_dec_eq(x_21, x_36); if (x_37 == 0) { lean_object* x_38; uint8_t x_39; -x_38 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__1; -x_39 = lean_string_dec_eq(x_19, x_38); +x_38 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__7; +x_39 = lean_string_dec_eq(x_21, x_38); if (x_39 == 0) { lean_object* x_40; uint8_t x_41; -x_40 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__8; -x_41 = lean_string_dec_eq(x_19, x_40); -lean_dec(x_19); +x_40 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__1; +x_41 = lean_string_dec_eq(x_21, x_40); if (x_41 == 0) { -lean_object* x_42; -lean_dec(x_18); -lean_dec(x_17); -x_42 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_42; +lean_object* x_42; uint8_t x_43; +x_42 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__8; +x_43 = lean_string_dec_eq(x_21, x_42); +lean_dec(x_21); +if (x_43 == 0) +{ +lean_object* x_44; +lean_dec(x_20); +lean_dec(x_19); +x_44 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_44; } else { -lean_object* x_43; uint8_t x_44; -x_43 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__9; -x_44 = lean_string_dec_eq(x_18, x_43); -if (x_44 == 0) -{ lean_object* x_45; uint8_t x_46; -x_45 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__10; -x_46 = lean_string_dec_eq(x_18, x_45); -lean_dec(x_18); +x_45 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__9; +x_46 = lean_string_dec_eq(x_20, x_45); if (x_46 == 0) { -lean_object* x_47; -lean_dec(x_17); -x_47 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_47; +lean_object* x_47; uint8_t x_48; +x_47 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__10; +x_48 = lean_string_dec_eq(x_20, x_47); +lean_dec(x_20); +if (x_48 == 0) +{ +lean_object* x_49; +lean_dec(x_19); +x_49 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_49; } else { -lean_object* x_48; lean_object* x_49; uint8_t x_50; -x_48 = lean_array_get_size(x_17); -x_49 = lean_unsigned_to_nat(3u); -x_50 = lean_nat_dec_eq(x_48, x_49); -lean_dec(x_48); -if (x_50 == 0) +lean_object* x_50; lean_object* x_51; uint8_t x_52; +x_50 = lean_array_get_size(x_19); +x_51 = lean_unsigned_to_nat(3u); +x_52 = lean_nat_dec_eq(x_50, x_51); +lean_dec(x_50); +if (x_52 == 0) { -lean_object* x_51; -lean_dec(x_17); -x_51 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_51; +lean_object* x_53; +lean_dec(x_19); +x_53 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_53; } else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_52 = lean_unsigned_to_nat(0u); -x_53 = lean_array_fget(x_17, x_52); -x_54 = lean_unsigned_to_nat(1u); -x_55 = lean_array_fget(x_17, x_54); -x_56 = lean_unsigned_to_nat(2u); -x_57 = lean_array_fget(x_17, x_56); -lean_dec(x_17); -if (lean_obj_tag(x_57) == 5) -{ -lean_object* x_58; -x_58 = lean_ctor_get(x_57, 0); -lean_inc(x_58); -if (lean_obj_tag(x_58) == 5) -{ -lean_object* x_59; -x_59 = lean_ctor_get(x_58, 0); -lean_inc(x_59); +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_54 = lean_unsigned_to_nat(0u); +x_55 = lean_array_fget(x_19, x_54); +x_56 = lean_unsigned_to_nat(1u); +x_57 = lean_array_fget(x_19, x_56); +x_58 = lean_unsigned_to_nat(2u); +x_59 = lean_array_fget(x_19, x_58); +lean_dec(x_19); if (lean_obj_tag(x_59) == 5) { lean_object* x_60; x_60 = lean_ctor_get(x_59, 0); lean_inc(x_60); -lean_dec(x_59); if (lean_obj_tag(x_60) == 5) { lean_object* x_61; x_61 = lean_ctor_get(x_60, 0); lean_inc(x_61); -lean_dec(x_60); -if (lean_obj_tag(x_61) == 4) +if (lean_obj_tag(x_61) == 5) { lean_object* x_62; x_62 = lean_ctor_get(x_61, 0); lean_inc(x_62); -if (lean_obj_tag(x_62) == 1) +lean_dec(x_61); +if (lean_obj_tag(x_62) == 5) { lean_object* x_63; x_63 = lean_ctor_get(x_62, 0); lean_inc(x_63); -if (lean_obj_tag(x_63) == 1) +lean_dec(x_62); +if (lean_obj_tag(x_63) == 4) { lean_object* x_64; x_64 = lean_ctor_get(x_63, 0); lean_inc(x_64); -if (lean_obj_tag(x_64) == 0) +if (lean_obj_tag(x_64) == 1) { -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_65 = lean_ctor_get(x_57, 1); +lean_object* x_65; +x_65 = lean_ctor_get(x_64, 0); lean_inc(x_65); -lean_dec(x_57); -x_66 = lean_ctor_get(x_58, 1); +if (lean_obj_tag(x_65) == 1) +{ +lean_object* x_66; +x_66 = lean_ctor_get(x_65, 0); lean_inc(x_66); -lean_dec(x_58); -x_67 = lean_ctor_get(x_61, 1); +if (lean_obj_tag(x_66) == 0) +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; uint8_t x_72; +x_67 = lean_ctor_get(x_59, 1); lean_inc(x_67); -lean_dec(x_61); -x_68 = lean_ctor_get(x_62, 1); +lean_dec(x_59); +x_68 = lean_ctor_get(x_60, 1); lean_inc(x_68); -lean_dec(x_62); +lean_dec(x_60); x_69 = lean_ctor_get(x_63, 1); lean_inc(x_69); lean_dec(x_63); -x_70 = lean_string_dec_eq(x_69, x_40); -lean_dec(x_69); -if (x_70 == 0) -{ -lean_object* x_71; -lean_dec(x_68); -lean_dec(x_67); -lean_dec(x_66); +x_70 = lean_ctor_get(x_64, 1); +lean_inc(x_70); +lean_dec(x_64); +x_71 = lean_ctor_get(x_65, 1); +lean_inc(x_71); lean_dec(x_65); -lean_dec(x_55); -lean_dec(x_53); -x_71 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_71; -} -else +x_72 = lean_string_dec_eq(x_71, x_42); +lean_dec(x_71); +if (x_72 == 0) { -lean_object* x_72; uint8_t x_73; -x_72 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__5; -x_73 = lean_string_dec_eq(x_68, x_72); +lean_object* x_73; +lean_dec(x_70); +lean_dec(x_69); lean_dec(x_68); -if (x_73 == 0) -{ -lean_object* x_74; lean_dec(x_67); -lean_dec(x_66); -lean_dec(x_65); +lean_dec(x_57); lean_dec(x_55); -lean_dec(x_53); -x_74 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_74; +x_73 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_73; } else { -if (lean_obj_tag(x_67) == 0) +lean_object* x_74; uint8_t x_75; +x_74 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__5; +x_75 = lean_string_dec_eq(x_70, x_74); +lean_dec(x_70); +if (x_75 == 0) { -lean_object* x_75; -lean_dec(x_66); -lean_dec(x_65); +lean_object* x_76; +lean_dec(x_69); +lean_dec(x_68); +lean_dec(x_67); +lean_dec(x_57); lean_dec(x_55); -lean_dec(x_53); -x_75 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_75; +x_76 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_76; } else { -lean_object* x_76; -x_76 = lean_ctor_get(x_67, 1); -lean_inc(x_76); -if (lean_obj_tag(x_76) == 0) +if (lean_obj_tag(x_69) == 0) { lean_object* x_77; +lean_dec(x_68); lean_dec(x_67); -lean_dec(x_66); -lean_dec(x_65); +lean_dec(x_57); lean_dec(x_55); -lean_dec(x_53); -x_77 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_77 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); return x_77; } else { lean_object* x_78; -x_78 = lean_ctor_get(x_76, 1); +x_78 = lean_ctor_get(x_69, 1); lean_inc(x_78); if (lean_obj_tag(x_78) == 0) { -uint8_t x_79; -x_79 = !lean_is_exclusive(x_67); -if (x_79 == 0) +lean_object* x_79; +lean_dec(x_69); +lean_dec(x_68); +lean_dec(x_67); +lean_dec(x_57); +lean_dec(x_55); +x_79 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_79; +} +else { -lean_object* x_80; uint8_t x_81; -x_80 = lean_ctor_get(x_67, 1); -lean_dec(x_80); -x_81 = !lean_is_exclusive(x_76); +lean_object* x_80; +x_80 = lean_ctor_get(x_78, 1); +lean_inc(x_80); +if (lean_obj_tag(x_80) == 0) +{ +uint8_t x_81; +x_81 = !lean_is_exclusive(x_69); if (x_81 == 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; -x_82 = lean_ctor_get(x_76, 1); +lean_object* x_82; uint8_t x_83; +x_82 = lean_ctor_get(x_69, 1); lean_dec(x_82); -x_83 = lean_box(0); -lean_ctor_set(x_76, 1, x_83); -x_84 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__12; -x_85 = l_Lean_Expr_const___override(x_84, x_67); -x_86 = l_Lean_mkApp4(x_85, x_53, x_66, x_55, x_65); -x_87 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_86, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_87; +x_83 = !lean_is_exclusive(x_78); +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; +x_84 = lean_ctor_get(x_78, 1); +lean_dec(x_84); +x_85 = lean_box(0); +lean_ctor_set(x_78, 1, x_85); +x_86 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__12; +x_87 = l_Lean_Expr_const___override(x_86, x_69); +x_88 = l_Lean_mkApp4(x_87, x_55, x_68, x_57, x_67); +x_89 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_88, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_89; } else { -lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; -x_88 = lean_ctor_get(x_76, 0); -lean_inc(x_88); -lean_dec(x_76); -x_89 = lean_box(0); -x_90 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_90, 0, x_88); -lean_ctor_set(x_90, 1, x_89); -lean_ctor_set(x_67, 1, x_90); -x_91 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__12; -x_92 = l_Lean_Expr_const___override(x_91, x_67); -x_93 = l_Lean_mkApp4(x_92, x_53, x_66, x_55, x_65); -x_94 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_93, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_94; +lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; +x_90 = lean_ctor_get(x_78, 0); +lean_inc(x_90); +lean_dec(x_78); +x_91 = lean_box(0); +x_92 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_91); +lean_ctor_set(x_69, 1, x_92); +x_93 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__12; +x_94 = l_Lean_Expr_const___override(x_93, x_69); +x_95 = l_Lean_mkApp4(x_94, x_55, x_68, x_57, x_67); +x_96 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_95, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_96; } } else { -lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; -x_95 = lean_ctor_get(x_67, 0); -lean_inc(x_95); -lean_dec(x_67); -x_96 = lean_ctor_get(x_76, 0); -lean_inc(x_96); -if (lean_is_exclusive(x_76)) { - lean_ctor_release(x_76, 0); - lean_ctor_release(x_76, 1); - x_97 = x_76; +lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_97 = lean_ctor_get(x_69, 0); +lean_inc(x_97); +lean_dec(x_69); +x_98 = lean_ctor_get(x_78, 0); +lean_inc(x_98); +if (lean_is_exclusive(x_78)) { + lean_ctor_release(x_78, 0); + lean_ctor_release(x_78, 1); + x_99 = x_78; } else { - lean_dec_ref(x_76); - x_97 = lean_box(0); + lean_dec_ref(x_78); + x_99 = lean_box(0); } -x_98 = lean_box(0); -if (lean_is_scalar(x_97)) { - x_99 = lean_alloc_ctor(1, 2, 0); +x_100 = lean_box(0); +if (lean_is_scalar(x_99)) { + x_101 = lean_alloc_ctor(1, 2, 0); } else { - x_99 = x_97; -} -lean_ctor_set(x_99, 0, x_96); -lean_ctor_set(x_99, 1, x_98); -x_100 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_100, 0, x_95); -lean_ctor_set(x_100, 1, x_99); -x_101 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__12; -x_102 = l_Lean_Expr_const___override(x_101, x_100); -x_103 = l_Lean_mkApp4(x_102, x_53, x_66, x_55, x_65); -x_104 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_103, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_104; + x_101 = x_99; +} +lean_ctor_set(x_101, 0, x_98); +lean_ctor_set(x_101, 1, x_100); +x_102 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_102, 0, x_97); +lean_ctor_set(x_102, 1, x_101); +x_103 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__12; +x_104 = l_Lean_Expr_const___override(x_103, x_102); +x_105 = l_Lean_mkApp4(x_104, x_55, x_68, x_57, x_67); +x_106 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_105, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_106; } } else { -lean_object* x_105; +lean_object* x_107; +lean_dec(x_80); lean_dec(x_78); -lean_dec(x_76); +lean_dec(x_69); +lean_dec(x_68); lean_dec(x_67); -lean_dec(x_66); -lean_dec(x_65); +lean_dec(x_57); lean_dec(x_55); -lean_dec(x_53); -x_105 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_105; +x_107 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_107; } } } @@ -9286,336 +9382,336 @@ return x_105; } else { -lean_object* x_106; +lean_object* x_108; +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_58); -lean_dec(x_57); -lean_dec(x_55); -lean_dec(x_53); -x_106 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_106; -} -} -else -{ -lean_object* x_107; -lean_dec(x_63); -lean_dec(x_62); -lean_dec(x_61); -lean_dec(x_58); -lean_dec(x_57); -lean_dec(x_55); -lean_dec(x_53); -x_107 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_107; -} -} -else -{ -lean_object* x_108; -lean_dec(x_62); -lean_dec(x_61); -lean_dec(x_58); +lean_dec(x_60); +lean_dec(x_59); lean_dec(x_57); lean_dec(x_55); -lean_dec(x_53); -x_108 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_108 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); return x_108; } } else { lean_object* x_109; -lean_dec(x_61); -lean_dec(x_58); +lean_dec(x_65); +lean_dec(x_64); +lean_dec(x_63); +lean_dec(x_60); +lean_dec(x_59); lean_dec(x_57); lean_dec(x_55); -lean_dec(x_53); -x_109 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_109 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); return x_109; } } else { lean_object* x_110; +lean_dec(x_64); +lean_dec(x_63); lean_dec(x_60); -lean_dec(x_58); +lean_dec(x_59); lean_dec(x_57); lean_dec(x_55); -lean_dec(x_53); -x_110 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_110 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); return x_110; } } else { lean_object* x_111; +lean_dec(x_63); +lean_dec(x_60); lean_dec(x_59); -lean_dec(x_58); lean_dec(x_57); lean_dec(x_55); -lean_dec(x_53); -x_111 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_111 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); return x_111; } } else { lean_object* x_112; -lean_dec(x_58); +lean_dec(x_62); +lean_dec(x_60); +lean_dec(x_59); lean_dec(x_57); lean_dec(x_55); -lean_dec(x_53); -x_112 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_112 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); return x_112; } } else { lean_object* x_113; +lean_dec(x_61); +lean_dec(x_60); +lean_dec(x_59); lean_dec(x_57); lean_dec(x_55); -lean_dec(x_53); -x_113 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_113 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); return x_113; } } +else +{ +lean_object* x_114; +lean_dec(x_60); +lean_dec(x_59); +lean_dec(x_57); +lean_dec(x_55); +x_114 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_114; } } else { -lean_object* x_114; lean_object* x_115; uint8_t x_116; -lean_dec(x_18); -x_114 = lean_array_get_size(x_17); -x_115 = lean_unsigned_to_nat(3u); -x_116 = lean_nat_dec_eq(x_114, x_115); -lean_dec(x_114); -if (x_116 == 0) -{ -lean_object* x_117; -lean_dec(x_17); -x_117 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_117; +lean_object* x_115; +lean_dec(x_59); +lean_dec(x_57); +lean_dec(x_55); +x_115 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_115; +} +} +} } else { -lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; -x_118 = lean_unsigned_to_nat(0u); -x_119 = lean_array_fget(x_17, x_118); -x_120 = lean_unsigned_to_nat(1u); -x_121 = lean_array_fget(x_17, x_120); -x_122 = lean_unsigned_to_nat(2u); -x_123 = lean_array_fget(x_17, x_122); -lean_dec(x_17); -if (lean_obj_tag(x_123) == 5) +lean_object* x_116; lean_object* x_117; uint8_t x_118; +lean_dec(x_20); +x_116 = lean_array_get_size(x_19); +x_117 = lean_unsigned_to_nat(3u); +x_118 = lean_nat_dec_eq(x_116, x_117); +lean_dec(x_116); +if (x_118 == 0) { -lean_object* x_124; -x_124 = lean_ctor_get(x_123, 0); -lean_inc(x_124); -if (lean_obj_tag(x_124) == 5) +lean_object* x_119; +lean_dec(x_19); +x_119 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_119; +} +else { -lean_object* x_125; -x_125 = lean_ctor_get(x_124, 0); -lean_inc(x_125); +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_120 = lean_unsigned_to_nat(0u); +x_121 = lean_array_fget(x_19, x_120); +x_122 = lean_unsigned_to_nat(1u); +x_123 = lean_array_fget(x_19, x_122); +x_124 = lean_unsigned_to_nat(2u); +x_125 = lean_array_fget(x_19, x_124); +lean_dec(x_19); if (lean_obj_tag(x_125) == 5) { lean_object* x_126; x_126 = lean_ctor_get(x_125, 0); lean_inc(x_126); -lean_dec(x_125); if (lean_obj_tag(x_126) == 5) { lean_object* x_127; x_127 = lean_ctor_get(x_126, 0); lean_inc(x_127); -lean_dec(x_126); -if (lean_obj_tag(x_127) == 4) +if (lean_obj_tag(x_127) == 5) { lean_object* x_128; x_128 = lean_ctor_get(x_127, 0); lean_inc(x_128); -if (lean_obj_tag(x_128) == 1) +lean_dec(x_127); +if (lean_obj_tag(x_128) == 5) { lean_object* x_129; x_129 = lean_ctor_get(x_128, 0); lean_inc(x_129); -if (lean_obj_tag(x_129) == 1) +lean_dec(x_128); +if (lean_obj_tag(x_129) == 4) { lean_object* x_130; x_130 = lean_ctor_get(x_129, 0); lean_inc(x_130); -if (lean_obj_tag(x_130) == 0) +if (lean_obj_tag(x_130) == 1) { -lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; uint8_t x_136; -x_131 = lean_ctor_get(x_123, 1); +lean_object* x_131; +x_131 = lean_ctor_get(x_130, 0); lean_inc(x_131); -lean_dec(x_123); -x_132 = lean_ctor_get(x_124, 1); +if (lean_obj_tag(x_131) == 1) +{ +lean_object* x_132; +x_132 = lean_ctor_get(x_131, 0); lean_inc(x_132); -lean_dec(x_124); -x_133 = lean_ctor_get(x_127, 1); +if (lean_obj_tag(x_132) == 0) +{ +lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; uint8_t x_138; +x_133 = lean_ctor_get(x_125, 1); lean_inc(x_133); -lean_dec(x_127); -x_134 = lean_ctor_get(x_128, 1); +lean_dec(x_125); +x_134 = lean_ctor_get(x_126, 1); lean_inc(x_134); -lean_dec(x_128); +lean_dec(x_126); x_135 = lean_ctor_get(x_129, 1); lean_inc(x_135); lean_dec(x_129); -x_136 = lean_string_dec_eq(x_135, x_40); -lean_dec(x_135); -if (x_136 == 0) -{ -lean_object* x_137; -lean_dec(x_134); -lean_dec(x_133); -lean_dec(x_132); +x_136 = lean_ctor_get(x_130, 1); +lean_inc(x_136); +lean_dec(x_130); +x_137 = lean_ctor_get(x_131, 1); +lean_inc(x_137); lean_dec(x_131); -lean_dec(x_121); -lean_dec(x_119); -x_137 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_137; -} -else +x_138 = lean_string_dec_eq(x_137, x_42); +lean_dec(x_137); +if (x_138 == 0) { -lean_object* x_138; uint8_t x_139; -x_138 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__5; -x_139 = lean_string_dec_eq(x_134, x_138); +lean_object* x_139; +lean_dec(x_136); +lean_dec(x_135); lean_dec(x_134); -if (x_139 == 0) -{ -lean_object* x_140; lean_dec(x_133); -lean_dec(x_132); -lean_dec(x_131); +lean_dec(x_123); lean_dec(x_121); -lean_dec(x_119); -x_140 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_140; +x_139 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_139; } else { -if (lean_obj_tag(x_133) == 0) +lean_object* x_140; uint8_t x_141; +x_140 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__5; +x_141 = lean_string_dec_eq(x_136, x_140); +lean_dec(x_136); +if (x_141 == 0) { -lean_object* x_141; -lean_dec(x_132); -lean_dec(x_131); +lean_object* x_142; +lean_dec(x_135); +lean_dec(x_134); +lean_dec(x_133); +lean_dec(x_123); lean_dec(x_121); -lean_dec(x_119); -x_141 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_141; +x_142 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_142; } else { -lean_object* x_142; -x_142 = lean_ctor_get(x_133, 1); -lean_inc(x_142); -if (lean_obj_tag(x_142) == 0) +if (lean_obj_tag(x_135) == 0) { lean_object* x_143; +lean_dec(x_134); lean_dec(x_133); -lean_dec(x_132); -lean_dec(x_131); +lean_dec(x_123); lean_dec(x_121); -lean_dec(x_119); -x_143 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_143 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); return x_143; } else { lean_object* x_144; -x_144 = lean_ctor_get(x_142, 1); +x_144 = lean_ctor_get(x_135, 1); lean_inc(x_144); if (lean_obj_tag(x_144) == 0) { -uint8_t x_145; -x_145 = !lean_is_exclusive(x_133); -if (x_145 == 0) +lean_object* x_145; +lean_dec(x_135); +lean_dec(x_134); +lean_dec(x_133); +lean_dec(x_123); +lean_dec(x_121); +x_145 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_145; +} +else +{ +lean_object* x_146; +x_146 = lean_ctor_get(x_144, 1); +lean_inc(x_146); +if (lean_obj_tag(x_146) == 0) { -lean_object* x_146; uint8_t x_147; -x_146 = lean_ctor_get(x_133, 1); -lean_dec(x_146); -x_147 = !lean_is_exclusive(x_142); +uint8_t x_147; +x_147 = !lean_is_exclusive(x_135); if (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; -x_148 = lean_ctor_get(x_142, 1); +lean_object* x_148; uint8_t x_149; +x_148 = lean_ctor_get(x_135, 1); lean_dec(x_148); -x_149 = lean_box(0); -lean_ctor_set(x_142, 1, x_149); -x_150 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__14; -x_151 = l_Lean_Expr_const___override(x_150, x_133); -x_152 = l_Lean_mkApp4(x_151, x_119, x_132, x_121, x_131); -x_153 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_152, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_153; +x_149 = !lean_is_exclusive(x_144); +if (x_149 == 0) +{ +lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; +x_150 = lean_ctor_get(x_144, 1); +lean_dec(x_150); +x_151 = lean_box(0); +lean_ctor_set(x_144, 1, x_151); +x_152 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__14; +x_153 = l_Lean_Expr_const___override(x_152, x_135); +x_154 = l_Lean_mkApp4(x_153, x_121, x_134, x_123, x_133); +x_155 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_154, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_155; } else { -lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; -x_154 = lean_ctor_get(x_142, 0); -lean_inc(x_154); -lean_dec(x_142); -x_155 = lean_box(0); -x_156 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_156, 0, x_154); -lean_ctor_set(x_156, 1, x_155); -lean_ctor_set(x_133, 1, x_156); -x_157 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__14; -x_158 = l_Lean_Expr_const___override(x_157, x_133); -x_159 = l_Lean_mkApp4(x_158, x_119, x_132, x_121, x_131); -x_160 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_159, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return 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; +x_156 = lean_ctor_get(x_144, 0); +lean_inc(x_156); +lean_dec(x_144); +x_157 = lean_box(0); +x_158 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_158, 0, x_156); +lean_ctor_set(x_158, 1, x_157); +lean_ctor_set(x_135, 1, x_158); +x_159 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__14; +x_160 = l_Lean_Expr_const___override(x_159, x_135); +x_161 = l_Lean_mkApp4(x_160, x_121, x_134, x_123, x_133); +x_162 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_161, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_162; } } else { -lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; -x_161 = lean_ctor_get(x_133, 0); -lean_inc(x_161); -lean_dec(x_133); -x_162 = lean_ctor_get(x_142, 0); -lean_inc(x_162); -if (lean_is_exclusive(x_142)) { - lean_ctor_release(x_142, 0); - lean_ctor_release(x_142, 1); - x_163 = x_142; +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; +x_163 = lean_ctor_get(x_135, 0); +lean_inc(x_163); +lean_dec(x_135); +x_164 = lean_ctor_get(x_144, 0); +lean_inc(x_164); +if (lean_is_exclusive(x_144)) { + lean_ctor_release(x_144, 0); + lean_ctor_release(x_144, 1); + x_165 = x_144; } else { - lean_dec_ref(x_142); - x_163 = lean_box(0); + lean_dec_ref(x_144); + x_165 = lean_box(0); } -x_164 = lean_box(0); -if (lean_is_scalar(x_163)) { - x_165 = lean_alloc_ctor(1, 2, 0); +x_166 = lean_box(0); +if (lean_is_scalar(x_165)) { + x_167 = lean_alloc_ctor(1, 2, 0); } else { - x_165 = x_163; -} -lean_ctor_set(x_165, 0, x_162); -lean_ctor_set(x_165, 1, x_164); -x_166 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_166, 0, x_161); -lean_ctor_set(x_166, 1, x_165); -x_167 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__14; -x_168 = l_Lean_Expr_const___override(x_167, x_166); -x_169 = l_Lean_mkApp4(x_168, x_119, x_132, x_121, x_131); -x_170 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_169, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_170; + x_167 = x_165; +} +lean_ctor_set(x_167, 0, x_164); +lean_ctor_set(x_167, 1, x_166); +x_168 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_168, 0, x_163); +lean_ctor_set(x_168, 1, x_167); +x_169 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__14; +x_170 = l_Lean_Expr_const___override(x_169, x_168); +x_171 = l_Lean_mkApp4(x_170, x_121, x_134, x_123, x_133); +x_172 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_171, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_172; } } else { -lean_object* x_171; +lean_object* x_173; +lean_dec(x_146); lean_dec(x_144); -lean_dec(x_142); +lean_dec(x_135); +lean_dec(x_134); lean_dec(x_133); -lean_dec(x_132); -lean_dec(x_131); +lean_dec(x_123); lean_dec(x_121); -lean_dec(x_119); -x_171 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_171; +x_173 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_173; } } } @@ -9624,101 +9720,101 @@ return x_171; } else { -lean_object* x_172; +lean_object* x_174; +lean_dec(x_132); +lean_dec(x_131); lean_dec(x_130); lean_dec(x_129); -lean_dec(x_128); -lean_dec(x_127); -lean_dec(x_124); +lean_dec(x_126); +lean_dec(x_125); lean_dec(x_123); lean_dec(x_121); -lean_dec(x_119); -x_172 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_172; +x_174 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_174; } } else { -lean_object* x_173; +lean_object* x_175; +lean_dec(x_131); +lean_dec(x_130); lean_dec(x_129); -lean_dec(x_128); -lean_dec(x_127); -lean_dec(x_124); +lean_dec(x_126); +lean_dec(x_125); lean_dec(x_123); lean_dec(x_121); -lean_dec(x_119); -x_173 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_173; +x_175 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_175; } } else { -lean_object* x_174; -lean_dec(x_128); -lean_dec(x_127); -lean_dec(x_124); +lean_object* x_176; +lean_dec(x_130); +lean_dec(x_129); +lean_dec(x_126); +lean_dec(x_125); lean_dec(x_123); lean_dec(x_121); -lean_dec(x_119); -x_174 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_174; +x_176 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_176; } } else { -lean_object* x_175; -lean_dec(x_127); -lean_dec(x_124); +lean_object* x_177; +lean_dec(x_129); +lean_dec(x_126); +lean_dec(x_125); lean_dec(x_123); lean_dec(x_121); -lean_dec(x_119); -x_175 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_175; +x_177 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_177; } } else { -lean_object* x_176; +lean_object* x_178; +lean_dec(x_128); lean_dec(x_126); -lean_dec(x_124); +lean_dec(x_125); lean_dec(x_123); lean_dec(x_121); -lean_dec(x_119); -x_176 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_176; +x_178 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_178; } } else { -lean_object* x_177; +lean_object* x_179; +lean_dec(x_127); +lean_dec(x_126); lean_dec(x_125); -lean_dec(x_124); lean_dec(x_123); lean_dec(x_121); -lean_dec(x_119); -x_177 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_177; +x_179 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_179; } } else { -lean_object* x_178; -lean_dec(x_124); +lean_object* x_180; +lean_dec(x_126); +lean_dec(x_125); lean_dec(x_123); lean_dec(x_121); -lean_dec(x_119); -x_178 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_178; +x_180 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_180; } } else { -lean_object* x_179; +lean_object* x_181; +lean_dec(x_125); lean_dec(x_123); lean_dec(x_121); -lean_dec(x_119); -x_179 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_179; +x_181 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_181; } } } @@ -9726,118 +9822,118 @@ return x_179; } else { -lean_object* x_180; uint8_t x_181; -lean_dec(x_19); -x_180 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__15; -x_181 = lean_string_dec_eq(x_18, x_180); -lean_dec(x_18); -if (x_181 == 0) +lean_object* x_182; uint8_t x_183; +lean_dec(x_21); +x_182 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__15; +x_183 = lean_string_dec_eq(x_20, x_182); +lean_dec(x_20); +if (x_183 == 0) { -lean_object* x_182; -lean_dec(x_17); -x_182 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_182; +lean_object* x_184; +lean_dec(x_19); +x_184 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_184; } else { -lean_object* x_183; lean_object* x_184; uint8_t x_185; -x_183 = lean_array_get_size(x_17); -x_184 = lean_unsigned_to_nat(3u); -x_185 = lean_nat_dec_eq(x_183, x_184); -lean_dec(x_183); -if (x_185 == 0) +lean_object* x_185; lean_object* x_186; uint8_t x_187; +x_185 = lean_array_get_size(x_19); +x_186 = lean_unsigned_to_nat(3u); +x_187 = lean_nat_dec_eq(x_185, x_186); +lean_dec(x_185); +if (x_187 == 0) { -lean_object* x_186; -lean_dec(x_17); -x_186 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_186; +lean_object* x_188; +lean_dec(x_19); +x_188 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_188; } else { -lean_object* x_187; lean_object* x_188; -x_187 = lean_unsigned_to_nat(0u); -x_188 = lean_array_fget(x_17, x_187); -if (lean_obj_tag(x_188) == 4) -{ -lean_object* x_189; -x_189 = lean_ctor_get(x_188, 0); -lean_inc(x_189); -if (lean_obj_tag(x_189) == 1) -{ -lean_object* x_190; -x_190 = lean_ctor_get(x_189, 0); -lean_inc(x_190); -if (lean_obj_tag(x_190) == 0) +lean_object* x_189; lean_object* x_190; +x_189 = lean_unsigned_to_nat(0u); +x_190 = lean_array_fget(x_19, x_189); +if (lean_obj_tag(x_190) == 4) { -lean_object* x_191; lean_object* x_192; lean_object* x_193; uint8_t x_194; -x_191 = lean_ctor_get(x_188, 1); +lean_object* x_191; +x_191 = lean_ctor_get(x_190, 0); lean_inc(x_191); -lean_dec(x_188); -x_192 = lean_ctor_get(x_189, 1); +if (lean_obj_tag(x_191) == 1) +{ +lean_object* x_192; +x_192 = lean_ctor_get(x_191, 0); lean_inc(x_192); -lean_dec(x_189); -x_193 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__9; -x_194 = lean_string_dec_eq(x_192, x_193); -lean_dec(x_192); -if (x_194 == 0) +if (lean_obj_tag(x_192) == 0) { -lean_object* x_195; +lean_object* x_193; lean_object* x_194; lean_object* x_195; uint8_t x_196; +x_193 = lean_ctor_get(x_190, 1); +lean_inc(x_193); +lean_dec(x_190); +x_194 = lean_ctor_get(x_191, 1); +lean_inc(x_194); lean_dec(x_191); -lean_dec(x_17); -x_195 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_195; +x_195 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__9; +x_196 = lean_string_dec_eq(x_194, x_195); +lean_dec(x_194); +if (x_196 == 0) +{ +lean_object* x_197; +lean_dec(x_193); +lean_dec(x_19); +x_197 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_197; } else { -if (lean_obj_tag(x_191) == 0) +if (lean_obj_tag(x_193) == 0) { -lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; -x_196 = lean_unsigned_to_nat(1u); -x_197 = lean_array_fget(x_17, x_196); -x_198 = lean_unsigned_to_nat(2u); -x_199 = lean_array_fget(x_17, x_198); -lean_dec(x_17); -x_200 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast(x_1, x_197, x_199, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_200; -} +lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; +x_198 = lean_unsigned_to_nat(1u); +x_199 = lean_array_fget(x_19, x_198); +x_200 = lean_unsigned_to_nat(2u); +x_201 = lean_array_fget(x_19, x_200); +lean_dec(x_19); +x_202 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast(x_1, x_199, x_201, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_202; +} else { -lean_object* x_201; -lean_dec(x_191); -lean_dec(x_17); -x_201 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_201; +lean_object* x_203; +lean_dec(x_193); +lean_dec(x_19); +x_203 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_203; } } } else { -lean_object* x_202; +lean_object* x_204; +lean_dec(x_192); +lean_dec(x_191); lean_dec(x_190); -lean_dec(x_189); -lean_dec(x_188); -lean_dec(x_17); -x_202 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_202; +lean_dec(x_19); +x_204 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_204; } } else { -lean_object* x_203; -lean_dec(x_189); -lean_dec(x_188); -lean_dec(x_17); -x_203 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_203; +lean_object* x_205; +lean_dec(x_191); +lean_dec(x_190); +lean_dec(x_19); +x_205 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_205; } } else { -lean_object* x_204; -lean_dec(x_188); -lean_dec(x_17); -x_204 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_204; +lean_object* x_206; +lean_dec(x_190); +lean_dec(x_19); +x_206 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_206; } } } @@ -9845,72 +9941,72 @@ return x_204; } else { -lean_object* x_205; uint8_t x_206; -lean_dec(x_19); -x_205 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__16; -x_206 = lean_string_dec_eq(x_18, x_205); -lean_dec(x_18); -if (x_206 == 0) +lean_object* x_207; uint8_t x_208; +lean_dec(x_21); +x_207 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__16; +x_208 = lean_string_dec_eq(x_20, x_207); +lean_dec(x_20); +if (x_208 == 0) { -lean_object* x_207; -lean_dec(x_17); -x_207 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_207; +lean_object* x_209; +lean_dec(x_19); +x_209 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_209; } else { -lean_object* x_208; lean_object* x_209; uint8_t x_210; -x_208 = lean_array_get_size(x_17); -x_209 = lean_unsigned_to_nat(6u); -x_210 = lean_nat_dec_eq(x_208, x_209); -lean_dec(x_208); -if (x_210 == 0) +lean_object* x_210; lean_object* x_211; uint8_t x_212; +x_210 = lean_array_get_size(x_19); +x_211 = lean_unsigned_to_nat(6u); +x_212 = lean_nat_dec_eq(x_210, x_211); +lean_dec(x_210); +if (x_212 == 0) { -lean_object* x_211; -lean_dec(x_17); -x_211 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_211; +lean_object* x_213; +lean_dec(x_19); +x_213 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_213; } else { -lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; -x_212 = lean_unsigned_to_nat(4u); -x_213 = lean_array_fget(x_17, x_212); -x_214 = lean_unsigned_to_nat(5u); -x_215 = lean_array_fget(x_17, x_214); -lean_dec(x_17); -x_216 = l_Lean_Elab_Tactic_Omega_succ_x3f(x_215); -if (lean_obj_tag(x_216) == 0) +lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; +x_214 = lean_unsigned_to_nat(4u); +x_215 = lean_array_fget(x_19, x_214); +x_216 = lean_unsigned_to_nat(5u); +x_217 = lean_array_fget(x_19, x_216); +lean_dec(x_19); +x_218 = l_Lean_Elab_Tactic_Omega_succ_x3f(x_217); +if (lean_obj_tag(x_218) == 0) { -lean_object* x_217; -lean_dec(x_213); -x_217 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_217; +lean_object* x_219; +lean_dec(x_215); +x_219 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_219; } else { -lean_object* x_218; lean_object* x_219; -x_218 = lean_ctor_get(x_216, 0); -lean_inc(x_218); -lean_dec(x_216); -lean_inc(x_213); -x_219 = l_Lean_Elab_Tactic_Omega_groundInt_x3f(x_213); -if (lean_obj_tag(x_219) == 0) -{ -lean_object* x_220; +lean_object* x_220; lean_object* x_221; +x_220 = lean_ctor_get(x_218, 0); +lean_inc(x_220); lean_dec(x_218); -lean_dec(x_213); -x_220 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_220; +lean_inc(x_215); +x_221 = l_Lean_Elab_Tactic_Omega_groundInt_x3f(x_215); +if (lean_obj_tag(x_221) == 0) +{ +lean_object* x_222; +lean_dec(x_220); +lean_dec(x_215); +x_222 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_222; } else { -lean_object* x_221; lean_object* x_222; lean_object* x_223; -lean_dec(x_219); -x_221 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__19; -x_222 = l_Lean_mkAppB(x_221, x_213, x_218); -x_223 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_222, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_223; +lean_object* x_223; lean_object* x_224; lean_object* x_225; +lean_dec(x_221); +x_223 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__19; +x_224 = l_Lean_mkAppB(x_223, x_215, x_220); +x_225 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_224, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_225; } } } @@ -9919,55 +10015,55 @@ return x_223; } else { -lean_object* x_224; uint8_t x_225; -lean_dec(x_19); -x_224 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__20; -x_225 = lean_string_dec_eq(x_18, x_224); -lean_dec(x_18); -if (x_225 == 0) +lean_object* x_226; uint8_t x_227; +lean_dec(x_21); +x_226 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__20; +x_227 = lean_string_dec_eq(x_20, x_226); +lean_dec(x_20); +if (x_227 == 0) { -lean_object* x_226; -lean_dec(x_17); -x_226 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_226; +lean_object* x_228; +lean_dec(x_19); +x_228 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_228; } else { -lean_object* x_227; lean_object* x_228; uint8_t x_229; -x_227 = lean_array_get_size(x_17); -x_228 = lean_unsigned_to_nat(4u); -x_229 = lean_nat_dec_eq(x_227, x_228); -lean_dec(x_227); -if (x_229 == 0) +lean_object* x_229; lean_object* x_230; uint8_t x_231; +x_229 = lean_array_get_size(x_19); +x_230 = lean_unsigned_to_nat(4u); +x_231 = lean_nat_dec_eq(x_229, x_230); +lean_dec(x_229); +if (x_231 == 0) { -lean_object* x_230; -lean_dec(x_17); -x_230 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_230; +lean_object* x_232; +lean_dec(x_19); +x_232 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_232; } else { -uint8_t x_231; -x_231 = lean_ctor_get_uint8(x_5, 3); -if (x_231 == 0) +uint8_t x_233; +x_233 = lean_ctor_get_uint8(x_5, 3); +if (x_233 == 0) { -lean_object* x_232; -lean_dec(x_17); -x_232 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_232; +lean_object* x_234; +lean_dec(x_19); +x_234 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_234; } else { -lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; -x_233 = lean_unsigned_to_nat(2u); -x_234 = lean_array_fget(x_17, x_233); -x_235 = lean_unsigned_to_nat(3u); -x_236 = lean_array_fget(x_17, x_235); -lean_dec(x_17); -x_237 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__23; -x_238 = l_Lean_mkAppB(x_237, x_234, x_236); -x_239 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_238, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_239; +lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; +x_235 = lean_unsigned_to_nat(2u); +x_236 = lean_array_fget(x_19, x_235); +x_237 = lean_unsigned_to_nat(3u); +x_238 = lean_array_fget(x_19, x_237); +lean_dec(x_19); +x_239 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__23; +x_240 = l_Lean_mkAppB(x_239, x_236, x_238); +x_241 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_240, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_241; } } } @@ -9975,55 +10071,55 @@ return x_239; } else { -lean_object* x_240; uint8_t x_241; -lean_dec(x_19); -x_240 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__24; -x_241 = lean_string_dec_eq(x_18, x_240); -lean_dec(x_18); -if (x_241 == 0) +lean_object* x_242; uint8_t x_243; +lean_dec(x_21); +x_242 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__24; +x_243 = lean_string_dec_eq(x_20, x_242); +lean_dec(x_20); +if (x_243 == 0) { -lean_object* x_242; -lean_dec(x_17); -x_242 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_242; +lean_object* x_244; +lean_dec(x_19); +x_244 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_244; } else { -lean_object* x_243; lean_object* x_244; uint8_t x_245; -x_243 = lean_array_get_size(x_17); -x_244 = lean_unsigned_to_nat(4u); -x_245 = lean_nat_dec_eq(x_243, x_244); -lean_dec(x_243); -if (x_245 == 0) +lean_object* x_245; lean_object* x_246; uint8_t x_247; +x_245 = lean_array_get_size(x_19); +x_246 = lean_unsigned_to_nat(4u); +x_247 = lean_nat_dec_eq(x_245, x_246); +lean_dec(x_245); +if (x_247 == 0) { -lean_object* x_246; -lean_dec(x_17); -x_246 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_246; +lean_object* x_248; +lean_dec(x_19); +x_248 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_248; } else { -uint8_t x_247; -x_247 = lean_ctor_get_uint8(x_5, 3); -if (x_247 == 0) +uint8_t x_249; +x_249 = lean_ctor_get_uint8(x_5, 3); +if (x_249 == 0) { -lean_object* x_248; -lean_dec(x_17); -x_248 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_248; +lean_object* x_250; +lean_dec(x_19); +x_250 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_250; } else { -lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; -x_249 = lean_unsigned_to_nat(2u); -x_250 = lean_array_fget(x_17, x_249); -x_251 = lean_unsigned_to_nat(3u); -x_252 = lean_array_fget(x_17, x_251); -lean_dec(x_17); -x_253 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__27; -x_254 = l_Lean_mkAppB(x_253, x_250, x_252); -x_255 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_254, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_255; +lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; +x_251 = lean_unsigned_to_nat(2u); +x_252 = lean_array_fget(x_19, x_251); +x_253 = lean_unsigned_to_nat(3u); +x_254 = lean_array_fget(x_19, x_253); +lean_dec(x_19); +x_255 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__27; +x_256 = l_Lean_mkAppB(x_255, x_252, x_254); +x_257 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_256, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_257; } } } @@ -10031,271 +10127,273 @@ return x_255; } else { -lean_object* x_256; uint8_t x_257; -lean_dec(x_19); -x_256 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__28; -x_257 = lean_string_dec_eq(x_18, x_256); -lean_dec(x_18); -if (x_257 == 0) +lean_object* x_258; uint8_t x_259; +lean_dec(x_21); +x_258 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__28; +x_259 = lean_string_dec_eq(x_20, x_258); +lean_dec(x_20); +if (x_259 == 0) { -lean_object* x_258; -lean_dec(x_17); -x_258 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_258; +lean_object* x_260; +lean_dec(x_19); +x_260 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_260; } else { -lean_object* x_259; lean_object* x_260; uint8_t x_261; -x_259 = lean_array_get_size(x_17); -x_260 = lean_unsigned_to_nat(6u); -x_261 = lean_nat_dec_eq(x_259, x_260); -lean_dec(x_259); -if (x_261 == 0) +lean_object* x_261; lean_object* x_262; uint8_t x_263; +x_261 = lean_array_get_size(x_19); +x_262 = lean_unsigned_to_nat(6u); +x_263 = lean_nat_dec_eq(x_261, x_262); +lean_dec(x_261); +if (x_263 == 0) { -lean_object* x_262; -lean_dec(x_17); -x_262 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_262; +lean_object* x_264; +lean_dec(x_19); +x_264 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_264; } else { -lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; -x_263 = lean_unsigned_to_nat(4u); -x_264 = lean_array_fget(x_17, x_263); -x_265 = lean_unsigned_to_nat(5u); -x_266 = lean_array_fget(x_17, x_265); -lean_dec(x_17); -x_267 = l_Lean_Elab_Tactic_Omega_groundInt_x3f(x_266); -if (lean_obj_tag(x_267) == 0) +lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; +x_265 = lean_unsigned_to_nat(4u); +x_266 = lean_array_fget(x_19, x_265); +x_267 = lean_unsigned_to_nat(5u); +x_268 = lean_array_fget(x_19, x_267); +lean_dec(x_19); +x_269 = l_Lean_Elab_Tactic_Omega_groundInt_x3f(x_268); +if (lean_obj_tag(x_269) == 0) { -lean_object* x_268; -lean_dec(x_264); -x_268 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_268; +lean_object* x_270; +lean_dec(x_266); +x_270 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_270; } else { -lean_object* x_269; lean_object* x_270; uint8_t x_271; -x_269 = lean_ctor_get(x_267, 0); -lean_inc(x_269); -lean_dec(x_267); -x_270 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__8; -x_271 = lean_int_dec_eq(x_269, x_270); -if (x_271 == 0) +lean_object* x_271; lean_object* x_272; uint8_t x_273; +x_271 = lean_ctor_get(x_269, 0); +lean_inc(x_271); +lean_dec(x_269); +x_272 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__8; +x_273 = lean_int_dec_eq(x_271, x_272); +if (x_273 == 0) { -uint8_t x_272; lean_object* x_273; lean_object* x_274; +uint8_t x_274; lean_object* x_275; lean_object* x_276; lean_dec(x_1); -x_272 = lean_int_dec_le(x_270, x_269); -x_273 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__3___closed__5; -lean_inc(x_264); -x_274 = lean_array_push(x_273, x_264); -if (x_272 == 0) -{ -lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; -x_275 = lean_int_neg(x_269); -x_276 = l_Int_toNat(x_275); -x_277 = l_Lean_instToExprInt_mkNat(x_276); -x_278 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__31; -x_279 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__11; -x_280 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__30; -lean_inc(x_277); -x_281 = l_Lean_mkApp3(x_278, x_279, x_280, x_277); -x_282 = lean_array_push(x_274, x_281); -x_283 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__29; +x_274 = lean_int_dec_le(x_272, x_271); +x_275 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__3___closed__5; +lean_inc(x_266); +x_276 = lean_array_push(x_275, x_266); +if (x_274 == 0) +{ +lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; +x_277 = lean_int_neg(x_271); +x_278 = l_Int_toNat(x_277); +x_279 = l_Lean_instToExprInt_mkNat(x_278); +x_280 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__31; +x_281 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__11; +x_282 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__30; +lean_inc(x_279); +x_283 = l_Lean_mkApp3(x_280, x_281, x_282, x_279); +x_284 = lean_array_push(x_276, x_283); +x_285 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__29; +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_284 = l_Lean_Meta_mkAppM(x_283, x_282, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_284) == 0) -{ -lean_object* x_285; lean_object* x_286; uint8_t x_287; -x_285 = lean_ctor_get(x_284, 0); -lean_inc(x_285); -x_286 = lean_ctor_get(x_284, 1); -lean_inc(x_286); -lean_dec(x_284); -x_287 = lean_int_dec_lt(x_269, x_270); -lean_dec(x_269); -if (x_287 == 0) +x_286 = l_Lean_Meta_mkAppM(x_285, x_284, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_286) == 0) { -lean_object* x_288; +lean_object* x_287; lean_object* x_288; uint8_t x_289; +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 = lean_int_dec_lt(x_271, x_272); +lean_dec(x_271); +if (x_289 == 0) +{ +lean_object* x_290; +lean_dec(x_279); lean_dec(x_277); -lean_dec(x_275); -lean_dec(x_264); -x_288 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_285, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_286); -return x_288; +lean_dec(x_266); +x_290 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_287, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_288); +return x_290; } else { -uint8_t x_289; -x_289 = lean_int_dec_le(x_270, x_275); -if (x_289 == 0) +uint8_t x_291; +x_291 = lean_int_dec_le(x_272, x_277); +if (x_291 == 0) { -lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; +lean_object* x_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_dec(x_279); +x_292 = lean_int_neg(x_277); lean_dec(x_277); -x_290 = lean_int_neg(x_275); -lean_dec(x_275); -x_291 = l_Int_toNat(x_290); -lean_dec(x_290); -x_292 = l_Lean_instToExprInt_mkNat(x_291); -x_293 = l_Lean_mkApp3(x_278, x_279, x_280, x_292); -x_294 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__34; -x_295 = l_Lean_mkAppB(x_294, x_264, x_293); -x_296 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_285, x_295, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_286); -return x_296; +x_293 = l_Int_toNat(x_292); +lean_dec(x_292); +x_294 = l_Lean_instToExprInt_mkNat(x_293); +x_295 = l_Lean_mkApp3(x_280, x_281, x_282, x_294); +x_296 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__34; +x_297 = l_Lean_mkAppB(x_296, x_266, x_295); +x_298 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_287, x_297, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_288); +return x_298; } else { -lean_object* x_297; lean_object* x_298; lean_object* x_299; -lean_dec(x_275); -x_297 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__34; -x_298 = l_Lean_mkAppB(x_297, x_264, x_277); -x_299 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_285, x_298, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_286); -return x_299; +lean_object* x_299; lean_object* x_300; lean_object* x_301; +lean_dec(x_277); +x_299 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__34; +x_300 = l_Lean_mkAppB(x_299, x_266, x_279); +x_301 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_287, x_300, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_288); +return x_301; } } } else { -uint8_t x_300; +uint8_t x_302; +lean_dec(x_279); lean_dec(x_277); -lean_dec(x_275); -lean_dec(x_269); -lean_dec(x_264); +lean_dec(x_271); +lean_dec(x_266); +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_300 = !lean_is_exclusive(x_284); -if (x_300 == 0) +x_302 = !lean_is_exclusive(x_286); +if (x_302 == 0) { -return x_284; +return x_286; } else { -lean_object* x_301; lean_object* x_302; lean_object* x_303; -x_301 = lean_ctor_get(x_284, 0); -x_302 = lean_ctor_get(x_284, 1); -lean_inc(x_302); -lean_inc(x_301); -lean_dec(x_284); -x_303 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_303, 0, x_301); -lean_ctor_set(x_303, 1, x_302); -return x_303; +lean_object* x_303; lean_object* x_304; lean_object* x_305; +x_303 = lean_ctor_get(x_286, 0); +x_304 = lean_ctor_get(x_286, 1); +lean_inc(x_304); +lean_inc(x_303); +lean_dec(x_286); +x_305 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_305, 0, x_303); +lean_ctor_set(x_305, 1, x_304); +return x_305; } } } else { -lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; -x_304 = l_Int_toNat(x_269); -x_305 = l_Lean_instToExprInt_mkNat(x_304); -x_306 = lean_array_push(x_274, x_305); -x_307 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__29; +lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; +x_306 = l_Int_toNat(x_271); +x_307 = l_Lean_instToExprInt_mkNat(x_306); +x_308 = lean_array_push(x_276, x_307); +x_309 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__29; +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_308 = l_Lean_Meta_mkAppM(x_307, x_306, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_308) == 0) -{ -lean_object* x_309; lean_object* x_310; uint8_t x_311; -x_309 = lean_ctor_get(x_308, 0); -lean_inc(x_309); -x_310 = lean_ctor_get(x_308, 1); -lean_inc(x_310); -lean_dec(x_308); -x_311 = lean_int_dec_lt(x_269, x_270); -if (x_311 == 0) -{ -lean_object* x_312; -lean_dec(x_269); -lean_dec(x_264); -x_312 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_309, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_310); -return x_312; +x_310 = l_Lean_Meta_mkAppM(x_309, x_308, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_310) == 0) +{ +lean_object* x_311; lean_object* x_312; uint8_t x_313; +x_311 = lean_ctor_get(x_310, 0); +lean_inc(x_311); +x_312 = lean_ctor_get(x_310, 1); +lean_inc(x_312); +lean_dec(x_310); +x_313 = lean_int_dec_lt(x_271, x_272); +if (x_313 == 0) +{ +lean_object* x_314; +lean_dec(x_271); +lean_dec(x_266); +x_314 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_311, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_312); +return x_314; } else { -lean_object* x_313; uint8_t x_314; -x_313 = lean_int_neg(x_269); -lean_dec(x_269); -x_314 = lean_int_dec_le(x_270, x_313); -if (x_314 == 0) +lean_object* x_315; uint8_t x_316; +x_315 = lean_int_neg(x_271); +lean_dec(x_271); +x_316 = lean_int_dec_le(x_272, x_315); +if (x_316 == 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; -x_315 = lean_int_neg(x_313); -lean_dec(x_313); -x_316 = l_Int_toNat(x_315); +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; +x_317 = lean_int_neg(x_315); lean_dec(x_315); -x_317 = l_Lean_instToExprInt_mkNat(x_316); -x_318 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__27; -x_319 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__11; -x_320 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__30; -x_321 = l_Lean_mkApp3(x_318, x_319, x_320, x_317); -x_322 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__35; -x_323 = l_Lean_mkAppB(x_322, x_264, x_321); -x_324 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_309, x_323, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_310); -return x_324; +x_318 = l_Int_toNat(x_317); +lean_dec(x_317); +x_319 = l_Lean_instToExprInt_mkNat(x_318); +x_320 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__27; +x_321 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__11; +x_322 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__30; +x_323 = l_Lean_mkApp3(x_320, x_321, x_322, x_319); +x_324 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__35; +x_325 = l_Lean_mkAppB(x_324, x_266, x_323); +x_326 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_311, x_325, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_312); +return x_326; } else { -lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; -x_325 = l_Int_toNat(x_313); -lean_dec(x_313); -x_326 = l_Lean_instToExprInt_mkNat(x_325); -x_327 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__35; -x_328 = l_Lean_mkAppB(x_327, x_264, x_326); -x_329 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_309, x_328, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_310); -return x_329; +lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; +x_327 = l_Int_toNat(x_315); +lean_dec(x_315); +x_328 = l_Lean_instToExprInt_mkNat(x_327); +x_329 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__35; +x_330 = l_Lean_mkAppB(x_329, x_266, x_328); +x_331 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_311, x_330, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_312); +return x_331; } } } else { -uint8_t x_330; -lean_dec(x_269); -lean_dec(x_264); +uint8_t x_332; +lean_dec(x_271); +lean_dec(x_266); +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_330 = !lean_is_exclusive(x_308); -if (x_330 == 0) +x_332 = !lean_is_exclusive(x_310); +if (x_332 == 0) { -return x_308; +return x_310; } else { -lean_object* x_331; lean_object* x_332; lean_object* x_333; -x_331 = lean_ctor_get(x_308, 0); -x_332 = lean_ctor_get(x_308, 1); -lean_inc(x_332); -lean_inc(x_331); -lean_dec(x_308); -x_333 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_333, 0, x_331); -lean_ctor_set(x_333, 1, x_332); -return x_333; +lean_object* x_333; lean_object* x_334; lean_object* x_335; +x_333 = lean_ctor_get(x_310, 0); +x_334 = lean_ctor_get(x_310, 1); +lean_inc(x_334); +lean_inc(x_333); +lean_dec(x_310); +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; } } } } else { -lean_object* x_334; lean_object* x_335; lean_object* x_336; -lean_dec(x_269); -x_334 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__38; -x_335 = l_Lean_Expr_app___override(x_334, x_264); -x_336 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_335, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_336; +lean_object* x_336; lean_object* x_337; lean_object* x_338; +lean_dec(x_271); +x_336 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__38; +x_337 = l_Lean_Expr_app___override(x_336, x_266); +x_338 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_337, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_338; } } } @@ -10304,182 +10402,184 @@ return x_336; } else { -lean_object* x_337; uint8_t x_338; -lean_dec(x_19); -x_337 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__39; -x_338 = lean_string_dec_eq(x_18, x_337); -lean_dec(x_18); -if (x_338 == 0) +lean_object* x_339; uint8_t x_340; +lean_dec(x_21); +x_339 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__39; +x_340 = lean_string_dec_eq(x_20, x_339); +lean_dec(x_20); +if (x_340 == 0) { -lean_object* x_339; -lean_dec(x_17); -x_339 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_339; +lean_object* x_341; +lean_dec(x_19); +x_341 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_341; } else { -lean_object* x_340; lean_object* x_341; uint8_t x_342; -x_340 = lean_array_get_size(x_17); -x_341 = lean_unsigned_to_nat(6u); -x_342 = lean_nat_dec_eq(x_340, x_341); -lean_dec(x_340); -if (x_342 == 0) +lean_object* x_342; lean_object* x_343; uint8_t x_344; +x_342 = lean_array_get_size(x_19); +x_343 = lean_unsigned_to_nat(6u); +x_344 = lean_nat_dec_eq(x_342, x_343); +lean_dec(x_342); +if (x_344 == 0) { -lean_object* x_343; -lean_dec(x_17); -x_343 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_343; +lean_object* x_345; +lean_dec(x_19); +x_345 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_345; } else { -lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; -x_344 = lean_unsigned_to_nat(4u); -x_345 = lean_array_fget(x_17, x_344); -x_346 = lean_unsigned_to_nat(5u); -x_347 = lean_array_fget(x_17, x_346); -lean_dec(x_17); -x_348 = l_Lean_Elab_Tactic_Omega_groundNat_x3f(x_347); -if (lean_obj_tag(x_348) == 0) +lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; +x_346 = lean_unsigned_to_nat(4u); +x_347 = lean_array_fget(x_19, x_346); +x_348 = lean_unsigned_to_nat(5u); +x_349 = lean_array_fget(x_19, x_348); +lean_dec(x_19); +x_350 = l_Lean_Elab_Tactic_Omega_groundNat_x3f(x_349); +if (lean_obj_tag(x_350) == 0) { -lean_object* x_349; -lean_dec(x_345); -x_349 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_349; +lean_object* x_351; +lean_dec(x_347); +x_351 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_351; } else { -lean_object* x_350; lean_object* x_351; lean_object* x_352; uint8_t x_353; lean_object* x_354; lean_object* x_355; +lean_object* x_352; lean_object* x_353; lean_object* x_354; uint8_t x_355; lean_object* x_356; lean_object* x_357; lean_dec(x_1); -x_350 = lean_ctor_get(x_348, 0); -lean_inc(x_350); -lean_dec(x_348); -x_351 = lean_nat_to_int(x_350); -x_352 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__8; -x_353 = lean_int_dec_le(x_352, x_351); -x_354 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__3___closed__5; -lean_inc(x_345); -x_355 = lean_array_push(x_354, x_345); -if (x_353 == 0) -{ -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; -x_356 = lean_int_neg(x_351); -lean_dec(x_351); -x_357 = l_Int_toNat(x_356); -lean_dec(x_356); -x_358 = l_Lean_instToExprInt_mkNat(x_357); -x_359 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__31; -x_360 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__11; -x_361 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__30; -x_362 = l_Lean_mkApp3(x_359, x_360, x_361, x_358); -lean_inc(x_362); -x_363 = lean_array_push(x_355, x_362); -x_364 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__40; +x_352 = lean_ctor_get(x_350, 0); +lean_inc(x_352); +lean_dec(x_350); +x_353 = lean_nat_to_int(x_352); +x_354 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__8; +x_355 = lean_int_dec_le(x_354, x_353); +x_356 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__3___closed__5; +lean_inc(x_347); +x_357 = lean_array_push(x_356, x_347); +if (x_355 == 0) +{ +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; +x_358 = lean_int_neg(x_353); +lean_dec(x_353); +x_359 = l_Int_toNat(x_358); +lean_dec(x_358); +x_360 = l_Lean_instToExprInt_mkNat(x_359); +x_361 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__31; +x_362 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__11; +x_363 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__30; +x_364 = l_Lean_mkApp3(x_361, x_362, x_363, x_360); +lean_inc(x_364); +x_365 = lean_array_push(x_357, x_364); +x_366 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__40; +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_365 = l_Lean_Meta_mkAppM(x_364, x_363, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_365) == 0) +x_367 = l_Lean_Meta_mkAppM(x_366, x_365, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_367) == 0) { -lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; -x_366 = lean_ctor_get(x_365, 0); -lean_inc(x_366); -x_367 = lean_ctor_get(x_365, 1); -lean_inc(x_367); -lean_dec(x_365); -x_368 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__43; -x_369 = l_Lean_mkAppB(x_368, x_345, x_362); -x_370 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_366, x_369, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_367); -return x_370; +lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; +x_368 = lean_ctor_get(x_367, 0); +lean_inc(x_368); +x_369 = lean_ctor_get(x_367, 1); +lean_inc(x_369); +lean_dec(x_367); +x_370 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__43; +x_371 = l_Lean_mkAppB(x_370, x_347, x_364); +x_372 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_368, x_371, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_369); +return x_372; } else { -uint8_t x_371; -lean_dec(x_362); -lean_dec(x_345); +uint8_t x_373; +lean_dec(x_364); +lean_dec(x_347); +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_371 = !lean_is_exclusive(x_365); -if (x_371 == 0) +x_373 = !lean_is_exclusive(x_367); +if (x_373 == 0) { -return x_365; +return x_367; } else { -lean_object* x_372; lean_object* x_373; lean_object* x_374; -x_372 = lean_ctor_get(x_365, 0); -x_373 = lean_ctor_get(x_365, 1); -lean_inc(x_373); -lean_inc(x_372); -lean_dec(x_365); -x_374 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_374, 0, x_372); -lean_ctor_set(x_374, 1, x_373); -return x_374; +lean_object* x_374; lean_object* x_375; lean_object* x_376; +x_374 = lean_ctor_get(x_367, 0); +x_375 = lean_ctor_get(x_367, 1); +lean_inc(x_375); +lean_inc(x_374); +lean_dec(x_367); +x_376 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_376, 0, x_374); +lean_ctor_set(x_376, 1, x_375); +return x_376; } } } else { -lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; -x_375 = l_Int_toNat(x_351); -lean_dec(x_351); -x_376 = l_Lean_instToExprInt_mkNat(x_375); -lean_inc(x_376); -x_377 = lean_array_push(x_355, x_376); -x_378 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__40; +lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; +x_377 = l_Int_toNat(x_353); +lean_dec(x_353); +x_378 = l_Lean_instToExprInt_mkNat(x_377); +lean_inc(x_378); +x_379 = lean_array_push(x_357, x_378); +x_380 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__40; +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_379 = l_Lean_Meta_mkAppM(x_378, x_377, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_379) == 0) +x_381 = l_Lean_Meta_mkAppM(x_380, x_379, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_381) == 0) { -lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; -x_380 = lean_ctor_get(x_379, 0); -lean_inc(x_380); -x_381 = lean_ctor_get(x_379, 1); -lean_inc(x_381); -lean_dec(x_379); -x_382 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__44; -x_383 = l_Lean_mkAppB(x_382, x_345, x_376); -x_384 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_380, x_383, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_381); -return x_384; +lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; +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_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__44; +x_385 = l_Lean_mkAppB(x_384, x_347, x_378); +x_386 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_382, x_385, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_383); +return x_386; } else { -uint8_t x_385; -lean_dec(x_376); -lean_dec(x_345); +uint8_t x_387; +lean_dec(x_378); +lean_dec(x_347); +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_385 = !lean_is_exclusive(x_379); -if (x_385 == 0) +x_387 = !lean_is_exclusive(x_381); +if (x_387 == 0) { -return x_379; +return x_381; } else { -lean_object* x_386; lean_object* x_387; lean_object* x_388; -x_386 = lean_ctor_get(x_379, 0); -x_387 = lean_ctor_get(x_379, 1); -lean_inc(x_387); -lean_inc(x_386); -lean_dec(x_379); -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_388; lean_object* x_389; lean_object* x_390; +x_388 = lean_ctor_get(x_381, 0); +x_389 = lean_ctor_get(x_381, 1); +lean_inc(x_389); +lean_inc(x_388); +lean_dec(x_381); +x_390 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_390, 0, x_388); +lean_ctor_set(x_390, 1, x_389); +return x_390; } } } @@ -10490,136 +10590,139 @@ return x_388; } else { -lean_object* x_389; uint8_t x_390; -lean_dec(x_19); -x_389 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__45; -x_390 = lean_string_dec_eq(x_18, x_389); -lean_dec(x_18); -if (x_390 == 0) +lean_object* x_391; uint8_t x_392; +lean_dec(x_21); +x_391 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__45; +x_392 = lean_string_dec_eq(x_20, x_391); +lean_dec(x_20); +if (x_392 == 0) { -lean_object* x_391; -lean_dec(x_17); -x_391 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_391; +lean_object* x_393; +lean_dec(x_19); +x_393 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_393; } else { -lean_object* x_392; lean_object* x_393; uint8_t x_394; -x_392 = lean_array_get_size(x_17); -x_393 = lean_unsigned_to_nat(6u); -x_394 = lean_nat_dec_eq(x_392, x_393); -lean_dec(x_392); -if (x_394 == 0) +lean_object* x_394; lean_object* x_395; uint8_t x_396; +x_394 = lean_array_get_size(x_19); +x_395 = lean_unsigned_to_nat(6u); +x_396 = lean_nat_dec_eq(x_394, x_395); +lean_dec(x_394); +if (x_396 == 0) { -lean_object* x_395; -lean_dec(x_17); -x_395 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_395; +lean_object* x_397; +lean_dec(x_19); +x_397 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_397; } else { -lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; -x_396 = lean_unsigned_to_nat(4u); -x_397 = lean_array_fget(x_17, x_396); -x_398 = lean_unsigned_to_nat(5u); -x_399 = lean_array_fget(x_17, x_398); -lean_dec(x_17); -x_400 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearCombo), 9, 1); -lean_closure_set(x_400, 0, x_397); -x_401 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__4), 10, 1); -lean_closure_set(x_401, 0, x_399); -x_402 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg), 10, 2); -lean_closure_set(x_402, 0, x_400); -lean_closure_set(x_402, 1, x_401); +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; +x_398 = lean_unsigned_to_nat(4u); +x_399 = lean_array_fget(x_19, x_398); +x_400 = lean_unsigned_to_nat(5u); +x_401 = lean_array_fget(x_19, x_400); +lean_dec(x_19); +x_402 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearCombo___boxed), 11, 1); +lean_closure_set(x_402, 0, x_399); +x_403 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__4___boxed), 12, 1); +lean_closure_set(x_403, 0, x_401); +x_404 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg___boxed), 12, 2); +lean_closure_set(x_404, 0, x_402); +lean_closure_set(x_404, 1, x_403); +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); lean_inc(x_3); -x_403 = l_Lean_Elab_Tactic_Omega_commitWhen___rarg(x_402, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_403) == 0) +x_405 = l_Lean_Elab_Tactic_Omega_commitWhen___rarg(x_404, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_405) == 0) { -lean_object* x_404; -x_404 = lean_ctor_get(x_403, 0); -lean_inc(x_404); -if (lean_obj_tag(x_404) == 0) +lean_object* x_406; +x_406 = lean_ctor_get(x_405, 0); +lean_inc(x_406); +if (lean_obj_tag(x_406) == 0) { -lean_object* x_405; lean_object* x_406; -x_405 = lean_ctor_get(x_403, 1); -lean_inc(x_405); -lean_dec(x_403); -x_406 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_405); -return x_406; +lean_object* x_407; lean_object* x_408; +x_407 = lean_ctor_get(x_405, 1); +lean_inc(x_407); +lean_dec(x_405); +x_408 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_407); +return x_408; } else { -uint8_t x_407; +uint8_t x_409; +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_1); -x_407 = !lean_is_exclusive(x_403); -if (x_407 == 0) +x_409 = !lean_is_exclusive(x_405); +if (x_409 == 0) { -lean_object* x_408; lean_object* x_409; -x_408 = lean_ctor_get(x_403, 0); -lean_dec(x_408); -x_409 = lean_ctor_get(x_404, 0); -lean_inc(x_409); -lean_dec(x_404); -lean_ctor_set(x_403, 0, x_409); -return x_403; +lean_object* x_410; lean_object* x_411; +x_410 = lean_ctor_get(x_405, 0); +lean_dec(x_410); +x_411 = lean_ctor_get(x_406, 0); +lean_inc(x_411); +lean_dec(x_406); +lean_ctor_set(x_405, 0, x_411); +return x_405; } else { -lean_object* x_410; lean_object* x_411; lean_object* x_412; -x_410 = lean_ctor_get(x_403, 1); -lean_inc(x_410); -lean_dec(x_403); -x_411 = lean_ctor_get(x_404, 0); -lean_inc(x_411); -lean_dec(x_404); -x_412 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_412, 0, x_411); -lean_ctor_set(x_412, 1, x_410); -return x_412; +lean_object* x_412; lean_object* x_413; lean_object* x_414; +x_412 = lean_ctor_get(x_405, 1); +lean_inc(x_412); +lean_dec(x_405); +x_413 = lean_ctor_get(x_406, 0); +lean_inc(x_413); +lean_dec(x_406); +x_414 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_414, 0, x_413); +lean_ctor_set(x_414, 1, x_412); +return x_414; } } } else { -uint8_t x_413; +uint8_t x_415; +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_1); -x_413 = !lean_is_exclusive(x_403); -if (x_413 == 0) +x_415 = !lean_is_exclusive(x_405); +if (x_415 == 0) { -return x_403; +return x_405; } else { -lean_object* x_414; lean_object* x_415; lean_object* x_416; -x_414 = lean_ctor_get(x_403, 0); -x_415 = lean_ctor_get(x_403, 1); -lean_inc(x_415); -lean_inc(x_414); -lean_dec(x_403); -x_416 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_416, 0, x_414); -lean_ctor_set(x_416, 1, x_415); -return x_416; +lean_object* x_416; lean_object* x_417; lean_object* x_418; +x_416 = lean_ctor_get(x_405, 0); +x_417 = lean_ctor_get(x_405, 1); +lean_inc(x_417); +lean_inc(x_416); +lean_dec(x_405); +x_418 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_418, 0, x_416); +lean_ctor_set(x_418, 1, x_417); +return x_418; } } } @@ -10628,221 +10731,221 @@ return x_416; } else { -lean_object* x_417; uint8_t x_418; -lean_dec(x_19); -x_417 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__23; -x_418 = lean_string_dec_eq(x_18, x_417); -lean_dec(x_18); -if (x_418 == 0) +lean_object* x_419; uint8_t x_420; +lean_dec(x_21); +x_419 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__23; +x_420 = lean_string_dec_eq(x_20, x_419); +lean_dec(x_20); +if (x_420 == 0) { -lean_object* x_419; -lean_dec(x_17); -x_419 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_419; +lean_object* x_421; +lean_dec(x_19); +x_421 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_421; } else { -lean_object* x_420; lean_object* x_421; uint8_t x_422; -x_420 = lean_array_get_size(x_17); -x_421 = lean_unsigned_to_nat(3u); -x_422 = lean_nat_dec_eq(x_420, x_421); -lean_dec(x_420); -if (x_422 == 0) +lean_object* x_422; lean_object* x_423; uint8_t x_424; +x_422 = lean_array_get_size(x_19); +x_423 = lean_unsigned_to_nat(3u); +x_424 = lean_nat_dec_eq(x_422, x_423); +lean_dec(x_422); +if (x_424 == 0) { -lean_object* x_423; -lean_dec(x_17); -x_423 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_423; +lean_object* x_425; +lean_dec(x_19); +x_425 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_425; } else { -lean_object* x_424; lean_object* x_425; lean_object* x_426; +lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_dec(x_1); -x_424 = lean_unsigned_to_nat(2u); -x_425 = lean_array_fget(x_17, x_424); -lean_dec(x_17); -x_426 = l_Lean_Elab_Tactic_Omega_asLinearCombo(x_425, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_426) == 0) -{ -lean_object* x_427; lean_object* x_428; uint8_t x_429; -x_427 = lean_ctor_get(x_426, 0); -lean_inc(x_427); -x_428 = lean_ctor_get(x_427, 1); -lean_inc(x_428); -x_429 = !lean_is_exclusive(x_426); -if (x_429 == 0) -{ -lean_object* x_430; uint8_t x_431; -x_430 = lean_ctor_get(x_426, 0); -lean_dec(x_430); -x_431 = !lean_is_exclusive(x_427); +x_426 = lean_unsigned_to_nat(2u); +x_427 = lean_array_fget(x_19, x_426); +lean_dec(x_19); +x_428 = l_Lean_Elab_Tactic_Omega_asLinearCombo(x_427, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_428) == 0) +{ +lean_object* x_429; lean_object* x_430; uint8_t x_431; +x_429 = lean_ctor_get(x_428, 0); +lean_inc(x_429); +x_430 = lean_ctor_get(x_429, 1); +lean_inc(x_430); +x_431 = !lean_is_exclusive(x_428); if (x_431 == 0) { -lean_object* x_432; lean_object* x_433; uint8_t x_434; -x_432 = lean_ctor_get(x_427, 0); -x_433 = lean_ctor_get(x_427, 1); -lean_dec(x_433); -x_434 = !lean_is_exclusive(x_428); -if (x_434 == 0) -{ -lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; -x_435 = lean_ctor_get(x_428, 0); -lean_inc(x_432); -x_436 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__5___boxed), 11, 2); -lean_closure_set(x_436, 0, x_432); -lean_closure_set(x_436, 1, x_435); -x_437 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__46; -x_438 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg), 10, 2); -lean_closure_set(x_438, 0, x_437); -lean_closure_set(x_438, 1, x_436); -x_439 = l_Lean_Omega_LinearCombo_neg(x_432); -lean_ctor_set(x_428, 0, x_438); -lean_ctor_set(x_427, 0, x_439); -return x_426; -} -else -{ -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; -x_440 = lean_ctor_get(x_428, 0); -x_441 = lean_ctor_get(x_428, 1); -lean_inc(x_441); -lean_inc(x_440); -lean_dec(x_428); -lean_inc(x_432); -x_442 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__5___boxed), 11, 2); -lean_closure_set(x_442, 0, x_432); -lean_closure_set(x_442, 1, x_440); -x_443 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__46; -x_444 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg), 10, 2); -lean_closure_set(x_444, 0, x_443); +lean_object* x_432; uint8_t x_433; +x_432 = lean_ctor_get(x_428, 0); +lean_dec(x_432); +x_433 = !lean_is_exclusive(x_429); +if (x_433 == 0) +{ +lean_object* x_434; lean_object* x_435; uint8_t x_436; +x_434 = lean_ctor_get(x_429, 0); +x_435 = lean_ctor_get(x_429, 1); +lean_dec(x_435); +x_436 = !lean_is_exclusive(x_430); +if (x_436 == 0) +{ +lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; +x_437 = lean_ctor_get(x_430, 0); +lean_inc(x_434); +x_438 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__5___boxed), 13, 2); +lean_closure_set(x_438, 0, x_434); +lean_closure_set(x_438, 1, x_437); +x_439 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__46; +x_440 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg___boxed), 12, 2); +lean_closure_set(x_440, 0, x_439); +lean_closure_set(x_440, 1, x_438); +x_441 = l_Lean_Omega_LinearCombo_neg(x_434); +lean_ctor_set(x_430, 0, x_440); +lean_ctor_set(x_429, 0, x_441); +return x_428; +} +else +{ +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; +x_442 = lean_ctor_get(x_430, 0); +x_443 = lean_ctor_get(x_430, 1); +lean_inc(x_443); +lean_inc(x_442); +lean_dec(x_430); +lean_inc(x_434); +x_444 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__5___boxed), 13, 2); +lean_closure_set(x_444, 0, x_434); lean_closure_set(x_444, 1, x_442); -x_445 = l_Lean_Omega_LinearCombo_neg(x_432); -x_446 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_446, 0, x_444); -lean_ctor_set(x_446, 1, x_441); -lean_ctor_set(x_427, 1, x_446); -lean_ctor_set(x_427, 0, x_445); -return x_426; +x_445 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__46; +x_446 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg___boxed), 12, 2); +lean_closure_set(x_446, 0, x_445); +lean_closure_set(x_446, 1, x_444); +x_447 = l_Lean_Omega_LinearCombo_neg(x_434); +x_448 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_448, 0, x_446); +lean_ctor_set(x_448, 1, x_443); +lean_ctor_set(x_429, 1, x_448); +lean_ctor_set(x_429, 0, x_447); +return x_428; } } else { -lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; -x_447 = lean_ctor_get(x_427, 0); -lean_inc(x_447); -lean_dec(x_427); -x_448 = lean_ctor_get(x_428, 0); -lean_inc(x_448); -x_449 = lean_ctor_get(x_428, 1); +lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; +x_449 = lean_ctor_get(x_429, 0); lean_inc(x_449); -if (lean_is_exclusive(x_428)) { - lean_ctor_release(x_428, 0); - lean_ctor_release(x_428, 1); - x_450 = x_428; +lean_dec(x_429); +x_450 = lean_ctor_get(x_430, 0); +lean_inc(x_450); +x_451 = lean_ctor_get(x_430, 1); +lean_inc(x_451); +if (lean_is_exclusive(x_430)) { + lean_ctor_release(x_430, 0); + lean_ctor_release(x_430, 1); + x_452 = x_430; } else { - lean_dec_ref(x_428); - x_450 = lean_box(0); + lean_dec_ref(x_430); + x_452 = lean_box(0); } -lean_inc(x_447); -x_451 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__5___boxed), 11, 2); -lean_closure_set(x_451, 0, x_447); -lean_closure_set(x_451, 1, x_448); -x_452 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__46; -x_453 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg), 10, 2); -lean_closure_set(x_453, 0, x_452); -lean_closure_set(x_453, 1, x_451); -x_454 = l_Lean_Omega_LinearCombo_neg(x_447); -if (lean_is_scalar(x_450)) { - x_455 = lean_alloc_ctor(0, 2, 0); +lean_inc(x_449); +x_453 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__5___boxed), 13, 2); +lean_closure_set(x_453, 0, x_449); +lean_closure_set(x_453, 1, x_450); +x_454 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__46; +x_455 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg___boxed), 12, 2); +lean_closure_set(x_455, 0, x_454); +lean_closure_set(x_455, 1, x_453); +x_456 = l_Lean_Omega_LinearCombo_neg(x_449); +if (lean_is_scalar(x_452)) { + x_457 = lean_alloc_ctor(0, 2, 0); } else { - x_455 = x_450; + x_457 = x_452; } -lean_ctor_set(x_455, 0, x_453); -lean_ctor_set(x_455, 1, x_449); -x_456 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_456, 0, x_454); -lean_ctor_set(x_456, 1, x_455); -lean_ctor_set(x_426, 0, x_456); -return x_426; +lean_ctor_set(x_457, 0, x_455); +lean_ctor_set(x_457, 1, x_451); +x_458 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_458, 0, x_456); +lean_ctor_set(x_458, 1, x_457); +lean_ctor_set(x_428, 0, x_458); +return x_428; } } else { -lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; -x_457 = lean_ctor_get(x_426, 1); -lean_inc(x_457); -lean_dec(x_426); -x_458 = lean_ctor_get(x_427, 0); -lean_inc(x_458); -if (lean_is_exclusive(x_427)) { - lean_ctor_release(x_427, 0); - lean_ctor_release(x_427, 1); - x_459 = x_427; +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; +x_459 = lean_ctor_get(x_428, 1); +lean_inc(x_459); +lean_dec(x_428); +x_460 = lean_ctor_get(x_429, 0); +lean_inc(x_460); +if (lean_is_exclusive(x_429)) { + lean_ctor_release(x_429, 0); + lean_ctor_release(x_429, 1); + x_461 = x_429; } else { - lean_dec_ref(x_427); - x_459 = lean_box(0); + lean_dec_ref(x_429); + x_461 = lean_box(0); } -x_460 = lean_ctor_get(x_428, 0); -lean_inc(x_460); -x_461 = lean_ctor_get(x_428, 1); -lean_inc(x_461); -if (lean_is_exclusive(x_428)) { - lean_ctor_release(x_428, 0); - lean_ctor_release(x_428, 1); - x_462 = x_428; +x_462 = lean_ctor_get(x_430, 0); +lean_inc(x_462); +x_463 = lean_ctor_get(x_430, 1); +lean_inc(x_463); +if (lean_is_exclusive(x_430)) { + lean_ctor_release(x_430, 0); + lean_ctor_release(x_430, 1); + x_464 = x_430; } else { - lean_dec_ref(x_428); - x_462 = lean_box(0); -} -lean_inc(x_458); -x_463 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__5___boxed), 11, 2); -lean_closure_set(x_463, 0, x_458); -lean_closure_set(x_463, 1, x_460); -x_464 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__46; -x_465 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg), 10, 2); -lean_closure_set(x_465, 0, x_464); -lean_closure_set(x_465, 1, x_463); -x_466 = l_Lean_Omega_LinearCombo_neg(x_458); -if (lean_is_scalar(x_462)) { - x_467 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_430); + x_464 = lean_box(0); +} +lean_inc(x_460); +x_465 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__5___boxed), 13, 2); +lean_closure_set(x_465, 0, x_460); +lean_closure_set(x_465, 1, x_462); +x_466 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__46; +x_467 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg___boxed), 12, 2); +lean_closure_set(x_467, 0, x_466); +lean_closure_set(x_467, 1, x_465); +x_468 = l_Lean_Omega_LinearCombo_neg(x_460); +if (lean_is_scalar(x_464)) { + x_469 = lean_alloc_ctor(0, 2, 0); } else { - x_467 = x_462; + x_469 = x_464; } -lean_ctor_set(x_467, 0, x_465); -lean_ctor_set(x_467, 1, x_461); -if (lean_is_scalar(x_459)) { - x_468 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_469, 0, x_467); +lean_ctor_set(x_469, 1, x_463); +if (lean_is_scalar(x_461)) { + x_470 = lean_alloc_ctor(0, 2, 0); } else { - x_468 = x_459; + x_470 = x_461; } -lean_ctor_set(x_468, 0, x_466); -lean_ctor_set(x_468, 1, 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_457); -return x_469; +lean_ctor_set(x_470, 0, x_468); +lean_ctor_set(x_470, 1, x_469); +x_471 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_471, 0, x_470); +lean_ctor_set(x_471, 1, x_459); +return x_471; } } else { -uint8_t x_470; -x_470 = !lean_is_exclusive(x_426); -if (x_470 == 0) +uint8_t x_472; +x_472 = !lean_is_exclusive(x_428); +if (x_472 == 0) { -return x_426; +return x_428; } else { -lean_object* x_471; lean_object* x_472; lean_object* x_473; -x_471 = lean_ctor_get(x_426, 0); -x_472 = lean_ctor_get(x_426, 1); -lean_inc(x_472); -lean_inc(x_471); -lean_dec(x_426); -x_473 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_473, 0, x_471); -lean_ctor_set(x_473, 1, x_472); -return x_473; +lean_object* x_473; lean_object* x_474; lean_object* x_475; +x_473 = lean_ctor_get(x_428, 0); +x_474 = lean_ctor_get(x_428, 1); +lean_inc(x_474); +lean_inc(x_473); +lean_dec(x_428); +x_475 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_475, 0, x_473); +lean_ctor_set(x_475, 1, x_474); +return x_475; } } } @@ -10851,503 +10954,505 @@ return x_473; } else { -lean_object* x_474; uint8_t x_475; -lean_dec(x_19); -x_474 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__47; -x_475 = lean_string_dec_eq(x_18, x_474); -lean_dec(x_18); -if (x_475 == 0) +lean_object* x_476; uint8_t x_477; +lean_dec(x_21); +x_476 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__47; +x_477 = lean_string_dec_eq(x_20, x_476); +lean_dec(x_20); +if (x_477 == 0) { -lean_object* x_476; -lean_dec(x_17); -x_476 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_476; +lean_object* x_478; +lean_dec(x_19); +x_478 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_478; } else { -lean_object* x_477; lean_object* x_478; uint8_t x_479; -x_477 = lean_array_get_size(x_17); -x_478 = lean_unsigned_to_nat(6u); -x_479 = lean_nat_dec_eq(x_477, x_478); -lean_dec(x_477); -if (x_479 == 0) +lean_object* x_479; lean_object* x_480; uint8_t x_481; +x_479 = lean_array_get_size(x_19); +x_480 = lean_unsigned_to_nat(6u); +x_481 = lean_nat_dec_eq(x_479, x_480); +lean_dec(x_479); +if (x_481 == 0) { -lean_object* x_480; -lean_dec(x_17); -x_480 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_480; +lean_object* x_482; +lean_dec(x_19); +x_482 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_482; } else { -lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; +lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_dec(x_1); -x_481 = lean_unsigned_to_nat(4u); -x_482 = lean_array_fget(x_17, x_481); -x_483 = lean_unsigned_to_nat(5u); -x_484 = lean_array_fget(x_17, x_483); -lean_dec(x_17); +x_483 = lean_unsigned_to_nat(4u); +x_484 = lean_array_fget(x_19, x_483); +x_485 = lean_unsigned_to_nat(5u); +x_486 = lean_array_fget(x_19, x_485); +lean_dec(x_19); +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_485 = l_Lean_Elab_Tactic_Omega_asLinearCombo(x_482, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_485) == 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; -x_486 = lean_ctor_get(x_485, 0); -lean_inc(x_486); -x_487 = lean_ctor_get(x_486, 1); -lean_inc(x_487); -x_488 = lean_ctor_get(x_485, 1); +x_487 = l_Lean_Elab_Tactic_Omega_asLinearCombo(x_484, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_487) == 0) +{ +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; +x_488 = lean_ctor_get(x_487, 0); lean_inc(x_488); -lean_dec(x_485); -x_489 = lean_ctor_get(x_486, 0); +x_489 = lean_ctor_get(x_488, 1); lean_inc(x_489); -lean_dec(x_486); -x_490 = lean_ctor_get(x_487, 0); +x_490 = lean_ctor_get(x_487, 1); lean_inc(x_490); -x_491 = lean_ctor_get(x_487, 1); -lean_inc(x_491); lean_dec(x_487); -x_492 = l_Lean_Elab_Tactic_Omega_asLinearCombo(x_484, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_488); -if (lean_obj_tag(x_492) == 0) -{ -lean_object* x_493; lean_object* x_494; uint8_t x_495; -x_493 = lean_ctor_get(x_492, 0); +x_491 = lean_ctor_get(x_488, 0); +lean_inc(x_491); +lean_dec(x_488); +x_492 = lean_ctor_get(x_489, 0); +lean_inc(x_492); +x_493 = lean_ctor_get(x_489, 1); lean_inc(x_493); -x_494 = lean_ctor_get(x_493, 1); -lean_inc(x_494); -x_495 = !lean_is_exclusive(x_492); -if (x_495 == 0) -{ -lean_object* x_496; uint8_t x_497; -x_496 = lean_ctor_get(x_492, 0); -lean_dec(x_496); -x_497 = !lean_is_exclusive(x_493); +lean_dec(x_489); +x_494 = l_Lean_Elab_Tactic_Omega_asLinearCombo(x_486, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_490); +if (lean_obj_tag(x_494) == 0) +{ +lean_object* x_495; lean_object* x_496; uint8_t x_497; +x_495 = lean_ctor_get(x_494, 0); +lean_inc(x_495); +x_496 = lean_ctor_get(x_495, 1); +lean_inc(x_496); +x_497 = !lean_is_exclusive(x_494); if (x_497 == 0) { -lean_object* x_498; lean_object* x_499; uint8_t x_500; -x_498 = lean_ctor_get(x_493, 0); -x_499 = lean_ctor_get(x_493, 1); -lean_dec(x_499); -x_500 = !lean_is_exclusive(x_494); -if (x_500 == 0) -{ -lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; uint8_t x_510; -x_501 = lean_ctor_get(x_494, 0); -x_502 = lean_ctor_get(x_494, 1); -lean_inc(x_498); -lean_inc(x_489); -x_503 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__6___boxed), 13, 4); -lean_closure_set(x_503, 0, x_489); -lean_closure_set(x_503, 1, x_498); -lean_closure_set(x_503, 2, x_490); -lean_closure_set(x_503, 3, x_501); -x_504 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__46; -x_505 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg), 10, 2); -lean_closure_set(x_505, 0, x_504); -lean_closure_set(x_505, 1, x_503); -x_506 = l_Lean_Omega_LinearCombo_sub(x_489, x_498); -x_507 = lean_ctor_get(x_502, 1); -lean_inc(x_507); -lean_dec(x_502); -x_508 = lean_array_get_size(x_507); -x_509 = lean_unsigned_to_nat(0u); -x_510 = lean_nat_dec_lt(x_509, x_508); -if (x_510 == 0) +lean_object* x_498; uint8_t x_499; +x_498 = lean_ctor_get(x_494, 0); +lean_dec(x_498); +x_499 = !lean_is_exclusive(x_495); +if (x_499 == 0) { -lean_dec(x_508); -lean_dec(x_507); -lean_ctor_set(x_494, 1, x_491); -lean_ctor_set(x_494, 0, x_505); -lean_ctor_set(x_493, 0, x_506); -return x_492; +lean_object* x_500; lean_object* x_501; uint8_t x_502; +x_500 = lean_ctor_get(x_495, 0); +x_501 = lean_ctor_get(x_495, 1); +lean_dec(x_501); +x_502 = !lean_is_exclusive(x_496); +if (x_502 == 0) +{ +lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; uint8_t x_512; +x_503 = lean_ctor_get(x_496, 0); +x_504 = lean_ctor_get(x_496, 1); +lean_inc(x_500); +lean_inc(x_491); +x_505 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__6___boxed), 15, 4); +lean_closure_set(x_505, 0, x_491); +lean_closure_set(x_505, 1, x_500); +lean_closure_set(x_505, 2, x_492); +lean_closure_set(x_505, 3, x_503); +x_506 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__46; +x_507 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg___boxed), 12, 2); +lean_closure_set(x_507, 0, x_506); +lean_closure_set(x_507, 1, x_505); +x_508 = l_Lean_Omega_LinearCombo_sub(x_491, x_500); +x_509 = lean_ctor_get(x_504, 1); +lean_inc(x_509); +lean_dec(x_504); +x_510 = lean_array_get_size(x_509); +x_511 = lean_unsigned_to_nat(0u); +x_512 = lean_nat_dec_lt(x_511, x_510); +if (x_512 == 0) +{ +lean_dec(x_510); +lean_dec(x_509); +lean_ctor_set(x_496, 1, x_493); +lean_ctor_set(x_496, 0, x_507); +lean_ctor_set(x_495, 0, x_508); +return x_494; } else { -uint8_t x_511; -x_511 = lean_nat_dec_le(x_508, x_508); -if (x_511 == 0) +uint8_t x_513; +x_513 = lean_nat_dec_le(x_510, x_510); +if (x_513 == 0) { -lean_dec(x_508); -lean_dec(x_507); -lean_ctor_set(x_494, 1, x_491); -lean_ctor_set(x_494, 0, x_505); -lean_ctor_set(x_493, 0, x_506); -return x_492; +lean_dec(x_510); +lean_dec(x_509); +lean_ctor_set(x_496, 1, x_493); +lean_ctor_set(x_496, 0, x_507); +lean_ctor_set(x_495, 0, x_508); +return x_494; } else { -size_t x_512; size_t x_513; lean_object* x_514; -x_512 = 0; -x_513 = lean_usize_of_nat(x_508); -lean_dec(x_508); -x_514 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Omega_asLinearComboImpl___spec__3(x_507, x_512, x_513, x_491); -lean_dec(x_507); -lean_ctor_set(x_494, 1, x_514); -lean_ctor_set(x_494, 0, x_505); -lean_ctor_set(x_493, 0, x_506); -return x_492; +size_t x_514; size_t x_515; lean_object* x_516; +x_514 = 0; +x_515 = lean_usize_of_nat(x_510); +lean_dec(x_510); +x_516 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Omega_asLinearComboImpl___spec__3(x_509, x_514, x_515, x_493); +lean_dec(x_509); +lean_ctor_set(x_496, 1, x_516); +lean_ctor_set(x_496, 0, x_507); +lean_ctor_set(x_495, 0, x_508); +return x_494; } } } else { -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; uint8_t x_524; -x_515 = lean_ctor_get(x_494, 0); -x_516 = lean_ctor_get(x_494, 1); -lean_inc(x_516); -lean_inc(x_515); -lean_dec(x_494); -lean_inc(x_498); -lean_inc(x_489); -x_517 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__6___boxed), 13, 4); -lean_closure_set(x_517, 0, x_489); -lean_closure_set(x_517, 1, x_498); -lean_closure_set(x_517, 2, x_490); -lean_closure_set(x_517, 3, x_515); -x_518 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__46; -x_519 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg), 10, 2); -lean_closure_set(x_519, 0, x_518); -lean_closure_set(x_519, 1, x_517); -x_520 = l_Lean_Omega_LinearCombo_sub(x_489, x_498); -x_521 = lean_ctor_get(x_516, 1); -lean_inc(x_521); -lean_dec(x_516); -x_522 = lean_array_get_size(x_521); -x_523 = lean_unsigned_to_nat(0u); -x_524 = lean_nat_dec_lt(x_523, x_522); -if (x_524 == 0) -{ -lean_object* x_525; -lean_dec(x_522); -lean_dec(x_521); -x_525 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_525, 0, x_519); -lean_ctor_set(x_525, 1, x_491); -lean_ctor_set(x_493, 1, x_525); -lean_ctor_set(x_493, 0, x_520); -return x_492; -} -else -{ -uint8_t x_526; -x_526 = lean_nat_dec_le(x_522, x_522); +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; uint8_t x_526; +x_517 = lean_ctor_get(x_496, 0); +x_518 = lean_ctor_get(x_496, 1); +lean_inc(x_518); +lean_inc(x_517); +lean_dec(x_496); +lean_inc(x_500); +lean_inc(x_491); +x_519 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__6___boxed), 15, 4); +lean_closure_set(x_519, 0, x_491); +lean_closure_set(x_519, 1, x_500); +lean_closure_set(x_519, 2, x_492); +lean_closure_set(x_519, 3, x_517); +x_520 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__46; +x_521 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg___boxed), 12, 2); +lean_closure_set(x_521, 0, x_520); +lean_closure_set(x_521, 1, x_519); +x_522 = l_Lean_Omega_LinearCombo_sub(x_491, x_500); +x_523 = lean_ctor_get(x_518, 1); +lean_inc(x_523); +lean_dec(x_518); +x_524 = lean_array_get_size(x_523); +x_525 = lean_unsigned_to_nat(0u); +x_526 = lean_nat_dec_lt(x_525, x_524); if (x_526 == 0) { lean_object* x_527; -lean_dec(x_522); -lean_dec(x_521); +lean_dec(x_524); +lean_dec(x_523); x_527 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_527, 0, x_519); -lean_ctor_set(x_527, 1, x_491); -lean_ctor_set(x_493, 1, x_527); -lean_ctor_set(x_493, 0, x_520); -return x_492; +lean_ctor_set(x_527, 0, x_521); +lean_ctor_set(x_527, 1, x_493); +lean_ctor_set(x_495, 1, x_527); +lean_ctor_set(x_495, 0, x_522); +return x_494; +} +else +{ +uint8_t x_528; +x_528 = lean_nat_dec_le(x_524, x_524); +if (x_528 == 0) +{ +lean_object* x_529; +lean_dec(x_524); +lean_dec(x_523); +x_529 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_529, 0, x_521); +lean_ctor_set(x_529, 1, x_493); +lean_ctor_set(x_495, 1, x_529); +lean_ctor_set(x_495, 0, x_522); +return x_494; } else { -size_t x_528; size_t x_529; lean_object* x_530; lean_object* x_531; -x_528 = 0; -x_529 = lean_usize_of_nat(x_522); -lean_dec(x_522); -x_530 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Omega_asLinearComboImpl___spec__3(x_521, x_528, x_529, x_491); -lean_dec(x_521); -x_531 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_531, 0, x_519); -lean_ctor_set(x_531, 1, x_530); -lean_ctor_set(x_493, 1, x_531); -lean_ctor_set(x_493, 0, x_520); -return x_492; +size_t x_530; size_t x_531; lean_object* x_532; lean_object* x_533; +x_530 = 0; +x_531 = lean_usize_of_nat(x_524); +lean_dec(x_524); +x_532 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Omega_asLinearComboImpl___spec__3(x_523, x_530, x_531, x_493); +lean_dec(x_523); +x_533 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_533, 0, x_521); +lean_ctor_set(x_533, 1, x_532); +lean_ctor_set(x_495, 1, x_533); +lean_ctor_set(x_495, 0, x_522); +return x_494; } } } } else { -lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; lean_object* x_541; lean_object* x_542; uint8_t x_543; -x_532 = lean_ctor_get(x_493, 0); -lean_inc(x_532); -lean_dec(x_493); -x_533 = lean_ctor_get(x_494, 0); -lean_inc(x_533); -x_534 = lean_ctor_get(x_494, 1); +lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; uint8_t x_545; +x_534 = lean_ctor_get(x_495, 0); lean_inc(x_534); -if (lean_is_exclusive(x_494)) { - lean_ctor_release(x_494, 0); - lean_ctor_release(x_494, 1); - x_535 = x_494; +lean_dec(x_495); +x_535 = lean_ctor_get(x_496, 0); +lean_inc(x_535); +x_536 = lean_ctor_get(x_496, 1); +lean_inc(x_536); +if (lean_is_exclusive(x_496)) { + lean_ctor_release(x_496, 0); + lean_ctor_release(x_496, 1); + x_537 = x_496; } else { - lean_dec_ref(x_494); - x_535 = lean_box(0); + lean_dec_ref(x_496); + x_537 = lean_box(0); } -lean_inc(x_532); -lean_inc(x_489); -x_536 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__6___boxed), 13, 4); -lean_closure_set(x_536, 0, x_489); -lean_closure_set(x_536, 1, x_532); -lean_closure_set(x_536, 2, x_490); -lean_closure_set(x_536, 3, x_533); -x_537 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__46; -x_538 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg), 10, 2); -lean_closure_set(x_538, 0, x_537); -lean_closure_set(x_538, 1, x_536); -x_539 = l_Lean_Omega_LinearCombo_sub(x_489, x_532); -x_540 = lean_ctor_get(x_534, 1); -lean_inc(x_540); -lean_dec(x_534); -x_541 = lean_array_get_size(x_540); -x_542 = lean_unsigned_to_nat(0u); -x_543 = lean_nat_dec_lt(x_542, x_541); -if (x_543 == 0) -{ -lean_object* x_544; lean_object* x_545; -lean_dec(x_541); -lean_dec(x_540); -if (lean_is_scalar(x_535)) { - x_544 = lean_alloc_ctor(0, 2, 0); +lean_inc(x_534); +lean_inc(x_491); +x_538 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__6___boxed), 15, 4); +lean_closure_set(x_538, 0, x_491); +lean_closure_set(x_538, 1, x_534); +lean_closure_set(x_538, 2, x_492); +lean_closure_set(x_538, 3, x_535); +x_539 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__46; +x_540 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg___boxed), 12, 2); +lean_closure_set(x_540, 0, x_539); +lean_closure_set(x_540, 1, x_538); +x_541 = l_Lean_Omega_LinearCombo_sub(x_491, x_534); +x_542 = lean_ctor_get(x_536, 1); +lean_inc(x_542); +lean_dec(x_536); +x_543 = lean_array_get_size(x_542); +x_544 = lean_unsigned_to_nat(0u); +x_545 = lean_nat_dec_lt(x_544, x_543); +if (x_545 == 0) +{ +lean_object* x_546; lean_object* x_547; +lean_dec(x_543); +lean_dec(x_542); +if (lean_is_scalar(x_537)) { + x_546 = lean_alloc_ctor(0, 2, 0); } else { - x_544 = x_535; + x_546 = x_537; } -lean_ctor_set(x_544, 0, x_538); -lean_ctor_set(x_544, 1, x_491); -x_545 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_545, 0, x_539); -lean_ctor_set(x_545, 1, x_544); -lean_ctor_set(x_492, 0, x_545); -return x_492; +lean_ctor_set(x_546, 0, x_540); +lean_ctor_set(x_546, 1, x_493); +x_547 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_547, 0, x_541); +lean_ctor_set(x_547, 1, x_546); +lean_ctor_set(x_494, 0, x_547); +return x_494; } else { -uint8_t x_546; -x_546 = lean_nat_dec_le(x_541, x_541); -if (x_546 == 0) +uint8_t x_548; +x_548 = lean_nat_dec_le(x_543, x_543); +if (x_548 == 0) { -lean_object* x_547; lean_object* x_548; -lean_dec(x_541); -lean_dec(x_540); -if (lean_is_scalar(x_535)) { - x_547 = lean_alloc_ctor(0, 2, 0); +lean_object* x_549; lean_object* x_550; +lean_dec(x_543); +lean_dec(x_542); +if (lean_is_scalar(x_537)) { + x_549 = lean_alloc_ctor(0, 2, 0); } else { - x_547 = x_535; + x_549 = x_537; } -lean_ctor_set(x_547, 0, x_538); -lean_ctor_set(x_547, 1, x_491); -x_548 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_548, 0, x_539); -lean_ctor_set(x_548, 1, x_547); -lean_ctor_set(x_492, 0, x_548); -return x_492; +lean_ctor_set(x_549, 0, x_540); +lean_ctor_set(x_549, 1, x_493); +x_550 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_550, 0, x_541); +lean_ctor_set(x_550, 1, x_549); +lean_ctor_set(x_494, 0, x_550); +return x_494; } else { -size_t x_549; size_t x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; -x_549 = 0; -x_550 = lean_usize_of_nat(x_541); -lean_dec(x_541); -x_551 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Omega_asLinearComboImpl___spec__3(x_540, x_549, x_550, x_491); -lean_dec(x_540); -if (lean_is_scalar(x_535)) { - x_552 = lean_alloc_ctor(0, 2, 0); +size_t x_551; size_t x_552; lean_object* x_553; lean_object* x_554; lean_object* x_555; +x_551 = 0; +x_552 = lean_usize_of_nat(x_543); +lean_dec(x_543); +x_553 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Omega_asLinearComboImpl___spec__3(x_542, x_551, x_552, x_493); +lean_dec(x_542); +if (lean_is_scalar(x_537)) { + x_554 = lean_alloc_ctor(0, 2, 0); } else { - x_552 = x_535; + x_554 = x_537; } -lean_ctor_set(x_552, 0, x_538); -lean_ctor_set(x_552, 1, x_551); -x_553 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_553, 0, x_539); -lean_ctor_set(x_553, 1, x_552); -lean_ctor_set(x_492, 0, x_553); -return x_492; +lean_ctor_set(x_554, 0, x_540); +lean_ctor_set(x_554, 1, x_553); +x_555 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_555, 0, x_541); +lean_ctor_set(x_555, 1, x_554); +lean_ctor_set(x_494, 0, x_555); +return x_494; } } } } else { -lean_object* x_554; lean_object* x_555; lean_object* x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; lean_object* x_564; lean_object* x_565; lean_object* x_566; uint8_t x_567; -x_554 = lean_ctor_get(x_492, 1); -lean_inc(x_554); -lean_dec(x_492); -x_555 = lean_ctor_get(x_493, 0); -lean_inc(x_555); -if (lean_is_exclusive(x_493)) { - lean_ctor_release(x_493, 0); - lean_ctor_release(x_493, 1); - x_556 = x_493; -} else { - lean_dec_ref(x_493); - x_556 = lean_box(0); -} -x_557 = lean_ctor_get(x_494, 0); +lean_object* x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; lean_object* x_564; lean_object* x_565; lean_object* x_566; lean_object* x_567; lean_object* x_568; uint8_t x_569; +x_556 = lean_ctor_get(x_494, 1); +lean_inc(x_556); +lean_dec(x_494); +x_557 = lean_ctor_get(x_495, 0); lean_inc(x_557); -x_558 = lean_ctor_get(x_494, 1); -lean_inc(x_558); -if (lean_is_exclusive(x_494)) { - lean_ctor_release(x_494, 0); - lean_ctor_release(x_494, 1); - x_559 = x_494; +if (lean_is_exclusive(x_495)) { + lean_ctor_release(x_495, 0); + lean_ctor_release(x_495, 1); + x_558 = x_495; } else { - lean_dec_ref(x_494); - x_559 = lean_box(0); + lean_dec_ref(x_495); + x_558 = lean_box(0); +} +x_559 = lean_ctor_get(x_496, 0); +lean_inc(x_559); +x_560 = lean_ctor_get(x_496, 1); +lean_inc(x_560); +if (lean_is_exclusive(x_496)) { + lean_ctor_release(x_496, 0); + lean_ctor_release(x_496, 1); + x_561 = x_496; +} else { + lean_dec_ref(x_496); + x_561 = lean_box(0); } -lean_inc(x_555); -lean_inc(x_489); -x_560 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__6___boxed), 13, 4); -lean_closure_set(x_560, 0, x_489); -lean_closure_set(x_560, 1, x_555); -lean_closure_set(x_560, 2, x_490); -lean_closure_set(x_560, 3, x_557); -x_561 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__46; -x_562 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg), 10, 2); -lean_closure_set(x_562, 0, x_561); -lean_closure_set(x_562, 1, x_560); -x_563 = l_Lean_Omega_LinearCombo_sub(x_489, x_555); -x_564 = lean_ctor_get(x_558, 1); -lean_inc(x_564); -lean_dec(x_558); -x_565 = lean_array_get_size(x_564); -x_566 = lean_unsigned_to_nat(0u); -x_567 = lean_nat_dec_lt(x_566, x_565); -if (x_567 == 0) -{ -lean_object* x_568; lean_object* x_569; lean_object* x_570; -lean_dec(x_565); -lean_dec(x_564); -if (lean_is_scalar(x_559)) { - x_568 = lean_alloc_ctor(0, 2, 0); +lean_inc(x_557); +lean_inc(x_491); +x_562 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__6___boxed), 15, 4); +lean_closure_set(x_562, 0, x_491); +lean_closure_set(x_562, 1, x_557); +lean_closure_set(x_562, 2, x_492); +lean_closure_set(x_562, 3, x_559); +x_563 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__46; +x_564 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg___boxed), 12, 2); +lean_closure_set(x_564, 0, x_563); +lean_closure_set(x_564, 1, x_562); +x_565 = l_Lean_Omega_LinearCombo_sub(x_491, x_557); +x_566 = lean_ctor_get(x_560, 1); +lean_inc(x_566); +lean_dec(x_560); +x_567 = lean_array_get_size(x_566); +x_568 = lean_unsigned_to_nat(0u); +x_569 = lean_nat_dec_lt(x_568, x_567); +if (x_569 == 0) +{ +lean_object* x_570; lean_object* x_571; lean_object* x_572; +lean_dec(x_567); +lean_dec(x_566); +if (lean_is_scalar(x_561)) { + x_570 = lean_alloc_ctor(0, 2, 0); } else { - x_568 = x_559; + x_570 = x_561; } -lean_ctor_set(x_568, 0, x_562); -lean_ctor_set(x_568, 1, x_491); -if (lean_is_scalar(x_556)) { - x_569 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_570, 0, x_564); +lean_ctor_set(x_570, 1, x_493); +if (lean_is_scalar(x_558)) { + x_571 = lean_alloc_ctor(0, 2, 0); } else { - x_569 = x_556; + x_571 = x_558; } -lean_ctor_set(x_569, 0, x_563); -lean_ctor_set(x_569, 1, x_568); -x_570 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_570, 0, x_569); -lean_ctor_set(x_570, 1, x_554); -return x_570; +lean_ctor_set(x_571, 0, x_565); +lean_ctor_set(x_571, 1, x_570); +x_572 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_572, 0, x_571); +lean_ctor_set(x_572, 1, x_556); +return x_572; } else { -uint8_t x_571; -x_571 = lean_nat_dec_le(x_565, x_565); -if (x_571 == 0) +uint8_t x_573; +x_573 = lean_nat_dec_le(x_567, x_567); +if (x_573 == 0) { -lean_object* x_572; lean_object* x_573; lean_object* x_574; -lean_dec(x_565); -lean_dec(x_564); -if (lean_is_scalar(x_559)) { - x_572 = lean_alloc_ctor(0, 2, 0); +lean_object* x_574; lean_object* x_575; lean_object* x_576; +lean_dec(x_567); +lean_dec(x_566); +if (lean_is_scalar(x_561)) { + x_574 = lean_alloc_ctor(0, 2, 0); } else { - x_572 = x_559; + x_574 = x_561; } -lean_ctor_set(x_572, 0, x_562); -lean_ctor_set(x_572, 1, x_491); -if (lean_is_scalar(x_556)) { - x_573 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_574, 0, x_564); +lean_ctor_set(x_574, 1, x_493); +if (lean_is_scalar(x_558)) { + x_575 = lean_alloc_ctor(0, 2, 0); } else { - x_573 = x_556; + x_575 = x_558; } -lean_ctor_set(x_573, 0, x_563); -lean_ctor_set(x_573, 1, x_572); -x_574 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_574, 0, x_573); -lean_ctor_set(x_574, 1, x_554); -return x_574; +lean_ctor_set(x_575, 0, x_565); +lean_ctor_set(x_575, 1, x_574); +x_576 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_576, 0, x_575); +lean_ctor_set(x_576, 1, x_556); +return x_576; } else { -size_t x_575; size_t x_576; lean_object* x_577; lean_object* x_578; lean_object* x_579; lean_object* x_580; -x_575 = 0; -x_576 = lean_usize_of_nat(x_565); -lean_dec(x_565); -x_577 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Omega_asLinearComboImpl___spec__3(x_564, x_575, x_576, x_491); -lean_dec(x_564); -if (lean_is_scalar(x_559)) { - x_578 = lean_alloc_ctor(0, 2, 0); +size_t x_577; size_t x_578; lean_object* x_579; lean_object* x_580; lean_object* x_581; lean_object* x_582; +x_577 = 0; +x_578 = lean_usize_of_nat(x_567); +lean_dec(x_567); +x_579 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Omega_asLinearComboImpl___spec__3(x_566, x_577, x_578, x_493); +lean_dec(x_566); +if (lean_is_scalar(x_561)) { + x_580 = lean_alloc_ctor(0, 2, 0); } else { - x_578 = x_559; + x_580 = x_561; } -lean_ctor_set(x_578, 0, x_562); -lean_ctor_set(x_578, 1, x_577); -if (lean_is_scalar(x_556)) { - x_579 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_580, 0, x_564); +lean_ctor_set(x_580, 1, x_579); +if (lean_is_scalar(x_558)) { + x_581 = lean_alloc_ctor(0, 2, 0); } else { - x_579 = x_556; + x_581 = x_558; } -lean_ctor_set(x_579, 0, x_563); -lean_ctor_set(x_579, 1, x_578); -x_580 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_580, 0, x_579); -lean_ctor_set(x_580, 1, x_554); -return x_580; +lean_ctor_set(x_581, 0, x_565); +lean_ctor_set(x_581, 1, x_580); +x_582 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_582, 0, x_581); +lean_ctor_set(x_582, 1, x_556); +return x_582; } } } } else { -uint8_t x_581; +uint8_t x_583; +lean_dec(x_493); +lean_dec(x_492); lean_dec(x_491); -lean_dec(x_490); -lean_dec(x_489); -x_581 = !lean_is_exclusive(x_492); -if (x_581 == 0) +x_583 = !lean_is_exclusive(x_494); +if (x_583 == 0) { -return x_492; +return x_494; } else { -lean_object* x_582; lean_object* x_583; lean_object* x_584; -x_582 = lean_ctor_get(x_492, 0); -x_583 = lean_ctor_get(x_492, 1); -lean_inc(x_583); -lean_inc(x_582); -lean_dec(x_492); -x_584 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_584, 0, x_582); -lean_ctor_set(x_584, 1, x_583); -return x_584; +lean_object* x_584; lean_object* x_585; lean_object* x_586; +x_584 = lean_ctor_get(x_494, 0); +x_585 = lean_ctor_get(x_494, 1); +lean_inc(x_585); +lean_inc(x_584); +lean_dec(x_494); +x_586 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_586, 0, x_584); +lean_ctor_set(x_586, 1, x_585); +return x_586; } } } else { -uint8_t x_585; -lean_dec(x_484); +uint8_t x_587; +lean_dec(x_486); +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_585 = !lean_is_exclusive(x_485); -if (x_585 == 0) +x_587 = !lean_is_exclusive(x_487); +if (x_587 == 0) { -return x_485; +return x_487; } else { -lean_object* x_586; lean_object* x_587; lean_object* x_588; -x_586 = lean_ctor_get(x_485, 0); -x_587 = lean_ctor_get(x_485, 1); -lean_inc(x_587); -lean_inc(x_586); -lean_dec(x_485); -x_588 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_588, 0, x_586); -lean_ctor_set(x_588, 1, x_587); -return x_588; +lean_object* x_588; lean_object* x_589; lean_object* x_590; +x_588 = lean_ctor_get(x_487, 0); +x_589 = lean_ctor_get(x_487, 1); +lean_inc(x_589); +lean_inc(x_588); +lean_dec(x_487); +x_590 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_590, 0, x_588); +lean_ctor_set(x_590, 1, x_589); +return x_590; } } } @@ -11356,503 +11461,505 @@ return x_588; } else { -lean_object* x_589; uint8_t x_590; -lean_dec(x_19); -x_589 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__3; -x_590 = lean_string_dec_eq(x_18, x_589); -lean_dec(x_18); -if (x_590 == 0) +lean_object* x_591; uint8_t x_592; +lean_dec(x_21); +x_591 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__3; +x_592 = lean_string_dec_eq(x_20, x_591); +lean_dec(x_20); +if (x_592 == 0) { -lean_object* x_591; -lean_dec(x_17); -x_591 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_591; +lean_object* x_593; +lean_dec(x_19); +x_593 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_593; } else { -lean_object* x_592; lean_object* x_593; uint8_t x_594; -x_592 = lean_array_get_size(x_17); -x_593 = lean_unsigned_to_nat(6u); -x_594 = lean_nat_dec_eq(x_592, x_593); -lean_dec(x_592); -if (x_594 == 0) +lean_object* x_594; lean_object* x_595; uint8_t x_596; +x_594 = lean_array_get_size(x_19); +x_595 = lean_unsigned_to_nat(6u); +x_596 = lean_nat_dec_eq(x_594, x_595); +lean_dec(x_594); +if (x_596 == 0) { -lean_object* x_595; -lean_dec(x_17); -x_595 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_595; +lean_object* x_597; +lean_dec(x_19); +x_597 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_597; } else { -lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; +lean_object* x_598; lean_object* x_599; lean_object* x_600; lean_object* x_601; lean_object* x_602; lean_dec(x_1); -x_596 = lean_unsigned_to_nat(4u); -x_597 = lean_array_fget(x_17, x_596); -x_598 = lean_unsigned_to_nat(5u); -x_599 = lean_array_fget(x_17, x_598); -lean_dec(x_17); +x_598 = lean_unsigned_to_nat(4u); +x_599 = lean_array_fget(x_19, x_598); +x_600 = lean_unsigned_to_nat(5u); +x_601 = lean_array_fget(x_19, x_600); +lean_dec(x_19); +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_600 = l_Lean_Elab_Tactic_Omega_asLinearCombo(x_597, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_600) == 0) -{ -lean_object* x_601; lean_object* x_602; lean_object* x_603; lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; -x_601 = lean_ctor_get(x_600, 0); -lean_inc(x_601); -x_602 = lean_ctor_get(x_601, 1); -lean_inc(x_602); -x_603 = lean_ctor_get(x_600, 1); +x_602 = l_Lean_Elab_Tactic_Omega_asLinearCombo(x_599, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_602) == 0) +{ +lean_object* x_603; lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; +x_603 = lean_ctor_get(x_602, 0); lean_inc(x_603); -lean_dec(x_600); -x_604 = lean_ctor_get(x_601, 0); +x_604 = lean_ctor_get(x_603, 1); lean_inc(x_604); -lean_dec(x_601); -x_605 = lean_ctor_get(x_602, 0); +x_605 = lean_ctor_get(x_602, 1); lean_inc(x_605); -x_606 = lean_ctor_get(x_602, 1); -lean_inc(x_606); lean_dec(x_602); -x_607 = l_Lean_Elab_Tactic_Omega_asLinearCombo(x_599, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_603); -if (lean_obj_tag(x_607) == 0) -{ -lean_object* x_608; lean_object* x_609; uint8_t x_610; -x_608 = lean_ctor_get(x_607, 0); +x_606 = lean_ctor_get(x_603, 0); +lean_inc(x_606); +lean_dec(x_603); +x_607 = lean_ctor_get(x_604, 0); +lean_inc(x_607); +x_608 = lean_ctor_get(x_604, 1); lean_inc(x_608); -x_609 = lean_ctor_get(x_608, 1); -lean_inc(x_609); -x_610 = !lean_is_exclusive(x_607); -if (x_610 == 0) +lean_dec(x_604); +x_609 = l_Lean_Elab_Tactic_Omega_asLinearCombo(x_601, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_605); +if (lean_obj_tag(x_609) == 0) { -lean_object* x_611; uint8_t x_612; -x_611 = lean_ctor_get(x_607, 0); -lean_dec(x_611); -x_612 = !lean_is_exclusive(x_608); +lean_object* x_610; lean_object* x_611; uint8_t x_612; +x_610 = lean_ctor_get(x_609, 0); +lean_inc(x_610); +x_611 = lean_ctor_get(x_610, 1); +lean_inc(x_611); +x_612 = !lean_is_exclusive(x_609); if (x_612 == 0) { -lean_object* x_613; lean_object* x_614; uint8_t x_615; -x_613 = lean_ctor_get(x_608, 0); -x_614 = lean_ctor_get(x_608, 1); -lean_dec(x_614); -x_615 = !lean_is_exclusive(x_609); -if (x_615 == 0) +lean_object* x_613; uint8_t x_614; +x_613 = lean_ctor_get(x_609, 0); +lean_dec(x_613); +x_614 = !lean_is_exclusive(x_610); +if (x_614 == 0) { -lean_object* x_616; lean_object* x_617; lean_object* x_618; lean_object* x_619; lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; lean_object* x_624; uint8_t x_625; -x_616 = lean_ctor_get(x_609, 0); -x_617 = lean_ctor_get(x_609, 1); -lean_inc(x_613); -lean_inc(x_604); -x_618 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__7___boxed), 13, 4); -lean_closure_set(x_618, 0, x_604); -lean_closure_set(x_618, 1, x_613); -lean_closure_set(x_618, 2, x_605); -lean_closure_set(x_618, 3, x_616); -x_619 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__46; -x_620 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg), 10, 2); -lean_closure_set(x_620, 0, x_619); -lean_closure_set(x_620, 1, x_618); -x_621 = l_Lean_Omega_LinearCombo_add(x_604, x_613); -x_622 = lean_ctor_get(x_617, 1); -lean_inc(x_622); -lean_dec(x_617); -x_623 = lean_array_get_size(x_622); -x_624 = lean_unsigned_to_nat(0u); -x_625 = lean_nat_dec_lt(x_624, x_623); -if (x_625 == 0) -{ -lean_dec(x_623); -lean_dec(x_622); -lean_ctor_set(x_609, 1, x_606); -lean_ctor_set(x_609, 0, x_620); -lean_ctor_set(x_608, 0, x_621); -return x_607; -} -else -{ -uint8_t x_626; -x_626 = lean_nat_dec_le(x_623, x_623); -if (x_626 == 0) -{ -lean_dec(x_623); -lean_dec(x_622); -lean_ctor_set(x_609, 1, x_606); -lean_ctor_set(x_609, 0, x_620); -lean_ctor_set(x_608, 0, x_621); -return x_607; -} -else -{ -size_t x_627; size_t x_628; lean_object* x_629; -x_627 = 0; -x_628 = lean_usize_of_nat(x_623); -lean_dec(x_623); -x_629 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Omega_asLinearComboImpl___spec__3(x_622, x_627, x_628, x_606); -lean_dec(x_622); -lean_ctor_set(x_609, 1, x_629); -lean_ctor_set(x_609, 0, x_620); -lean_ctor_set(x_608, 0, x_621); -return x_607; -} -} -} -else -{ -lean_object* x_630; lean_object* x_631; lean_object* x_632; lean_object* x_633; lean_object* x_634; lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; uint8_t x_639; -x_630 = lean_ctor_get(x_609, 0); -x_631 = lean_ctor_get(x_609, 1); -lean_inc(x_631); -lean_inc(x_630); -lean_dec(x_609); -lean_inc(x_613); -lean_inc(x_604); -x_632 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__7___boxed), 13, 4); -lean_closure_set(x_632, 0, x_604); -lean_closure_set(x_632, 1, x_613); -lean_closure_set(x_632, 2, x_605); -lean_closure_set(x_632, 3, x_630); -x_633 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__46; -x_634 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg), 10, 2); -lean_closure_set(x_634, 0, x_633); -lean_closure_set(x_634, 1, x_632); -x_635 = l_Lean_Omega_LinearCombo_add(x_604, x_613); -x_636 = lean_ctor_get(x_631, 1); -lean_inc(x_636); -lean_dec(x_631); -x_637 = lean_array_get_size(x_636); -x_638 = lean_unsigned_to_nat(0u); -x_639 = lean_nat_dec_lt(x_638, x_637); -if (x_639 == 0) -{ -lean_object* x_640; -lean_dec(x_637); -lean_dec(x_636); -x_640 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_640, 0, x_634); -lean_ctor_set(x_640, 1, x_606); -lean_ctor_set(x_608, 1, x_640); -lean_ctor_set(x_608, 0, x_635); -return x_607; -} -else -{ -uint8_t x_641; -x_641 = lean_nat_dec_le(x_637, x_637); +lean_object* x_615; lean_object* x_616; uint8_t x_617; +x_615 = lean_ctor_get(x_610, 0); +x_616 = lean_ctor_get(x_610, 1); +lean_dec(x_616); +x_617 = !lean_is_exclusive(x_611); +if (x_617 == 0) +{ +lean_object* x_618; lean_object* x_619; lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; lean_object* x_624; lean_object* x_625; lean_object* x_626; uint8_t x_627; +x_618 = lean_ctor_get(x_611, 0); +x_619 = lean_ctor_get(x_611, 1); +lean_inc(x_615); +lean_inc(x_606); +x_620 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__7___boxed), 15, 4); +lean_closure_set(x_620, 0, x_606); +lean_closure_set(x_620, 1, x_615); +lean_closure_set(x_620, 2, x_607); +lean_closure_set(x_620, 3, x_618); +x_621 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__46; +x_622 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg___boxed), 12, 2); +lean_closure_set(x_622, 0, x_621); +lean_closure_set(x_622, 1, x_620); +x_623 = l_Lean_Omega_LinearCombo_add(x_606, x_615); +x_624 = lean_ctor_get(x_619, 1); +lean_inc(x_624); +lean_dec(x_619); +x_625 = lean_array_get_size(x_624); +x_626 = lean_unsigned_to_nat(0u); +x_627 = lean_nat_dec_lt(x_626, x_625); +if (x_627 == 0) +{ +lean_dec(x_625); +lean_dec(x_624); +lean_ctor_set(x_611, 1, x_608); +lean_ctor_set(x_611, 0, x_622); +lean_ctor_set(x_610, 0, x_623); +return x_609; +} +else +{ +uint8_t x_628; +x_628 = lean_nat_dec_le(x_625, x_625); +if (x_628 == 0) +{ +lean_dec(x_625); +lean_dec(x_624); +lean_ctor_set(x_611, 1, x_608); +lean_ctor_set(x_611, 0, x_622); +lean_ctor_set(x_610, 0, x_623); +return x_609; +} +else +{ +size_t x_629; size_t x_630; lean_object* x_631; +x_629 = 0; +x_630 = lean_usize_of_nat(x_625); +lean_dec(x_625); +x_631 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Omega_asLinearComboImpl___spec__3(x_624, x_629, x_630, x_608); +lean_dec(x_624); +lean_ctor_set(x_611, 1, x_631); +lean_ctor_set(x_611, 0, x_622); +lean_ctor_set(x_610, 0, x_623); +return x_609; +} +} +} +else +{ +lean_object* x_632; lean_object* x_633; lean_object* x_634; lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; lean_object* x_639; lean_object* x_640; uint8_t x_641; +x_632 = lean_ctor_get(x_611, 0); +x_633 = lean_ctor_get(x_611, 1); +lean_inc(x_633); +lean_inc(x_632); +lean_dec(x_611); +lean_inc(x_615); +lean_inc(x_606); +x_634 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__7___boxed), 15, 4); +lean_closure_set(x_634, 0, x_606); +lean_closure_set(x_634, 1, x_615); +lean_closure_set(x_634, 2, x_607); +lean_closure_set(x_634, 3, x_632); +x_635 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__46; +x_636 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg___boxed), 12, 2); +lean_closure_set(x_636, 0, x_635); +lean_closure_set(x_636, 1, x_634); +x_637 = l_Lean_Omega_LinearCombo_add(x_606, x_615); +x_638 = lean_ctor_get(x_633, 1); +lean_inc(x_638); +lean_dec(x_633); +x_639 = lean_array_get_size(x_638); +x_640 = lean_unsigned_to_nat(0u); +x_641 = lean_nat_dec_lt(x_640, x_639); if (x_641 == 0) { lean_object* x_642; -lean_dec(x_637); -lean_dec(x_636); +lean_dec(x_639); +lean_dec(x_638); x_642 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_642, 0, x_634); -lean_ctor_set(x_642, 1, x_606); -lean_ctor_set(x_608, 1, x_642); -lean_ctor_set(x_608, 0, x_635); -return x_607; +lean_ctor_set(x_642, 0, x_636); +lean_ctor_set(x_642, 1, x_608); +lean_ctor_set(x_610, 1, x_642); +lean_ctor_set(x_610, 0, x_637); +return x_609; +} +else +{ +uint8_t x_643; +x_643 = lean_nat_dec_le(x_639, x_639); +if (x_643 == 0) +{ +lean_object* x_644; +lean_dec(x_639); +lean_dec(x_638); +x_644 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_644, 0, x_636); +lean_ctor_set(x_644, 1, x_608); +lean_ctor_set(x_610, 1, x_644); +lean_ctor_set(x_610, 0, x_637); +return x_609; } else { -size_t x_643; size_t x_644; lean_object* x_645; lean_object* x_646; -x_643 = 0; -x_644 = lean_usize_of_nat(x_637); -lean_dec(x_637); -x_645 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Omega_asLinearComboImpl___spec__3(x_636, x_643, x_644, x_606); -lean_dec(x_636); -x_646 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_646, 0, x_634); -lean_ctor_set(x_646, 1, x_645); -lean_ctor_set(x_608, 1, x_646); -lean_ctor_set(x_608, 0, x_635); -return x_607; +size_t x_645; size_t x_646; lean_object* x_647; lean_object* x_648; +x_645 = 0; +x_646 = lean_usize_of_nat(x_639); +lean_dec(x_639); +x_647 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Omega_asLinearComboImpl___spec__3(x_638, x_645, x_646, x_608); +lean_dec(x_638); +x_648 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_648, 0, x_636); +lean_ctor_set(x_648, 1, x_647); +lean_ctor_set(x_610, 1, x_648); +lean_ctor_set(x_610, 0, x_637); +return x_609; } } } } else { -lean_object* x_647; lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; lean_object* x_653; lean_object* x_654; lean_object* x_655; lean_object* x_656; lean_object* x_657; uint8_t x_658; -x_647 = lean_ctor_get(x_608, 0); -lean_inc(x_647); -lean_dec(x_608); -x_648 = lean_ctor_get(x_609, 0); -lean_inc(x_648); -x_649 = lean_ctor_get(x_609, 1); +lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; lean_object* x_653; lean_object* x_654; lean_object* x_655; lean_object* x_656; lean_object* x_657; lean_object* x_658; lean_object* x_659; uint8_t x_660; +x_649 = lean_ctor_get(x_610, 0); lean_inc(x_649); -if (lean_is_exclusive(x_609)) { - lean_ctor_release(x_609, 0); - lean_ctor_release(x_609, 1); - x_650 = x_609; +lean_dec(x_610); +x_650 = lean_ctor_get(x_611, 0); +lean_inc(x_650); +x_651 = lean_ctor_get(x_611, 1); +lean_inc(x_651); +if (lean_is_exclusive(x_611)) { + lean_ctor_release(x_611, 0); + lean_ctor_release(x_611, 1); + x_652 = x_611; } else { - lean_dec_ref(x_609); - x_650 = lean_box(0); + lean_dec_ref(x_611); + x_652 = lean_box(0); } -lean_inc(x_647); -lean_inc(x_604); -x_651 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__7___boxed), 13, 4); -lean_closure_set(x_651, 0, x_604); -lean_closure_set(x_651, 1, x_647); -lean_closure_set(x_651, 2, x_605); -lean_closure_set(x_651, 3, x_648); -x_652 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__46; -x_653 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg), 10, 2); -lean_closure_set(x_653, 0, x_652); -lean_closure_set(x_653, 1, x_651); -x_654 = l_Lean_Omega_LinearCombo_add(x_604, x_647); -x_655 = lean_ctor_get(x_649, 1); -lean_inc(x_655); -lean_dec(x_649); -x_656 = lean_array_get_size(x_655); -x_657 = lean_unsigned_to_nat(0u); -x_658 = lean_nat_dec_lt(x_657, x_656); -if (x_658 == 0) -{ -lean_object* x_659; lean_object* x_660; -lean_dec(x_656); -lean_dec(x_655); -if (lean_is_scalar(x_650)) { - x_659 = lean_alloc_ctor(0, 2, 0); +lean_inc(x_649); +lean_inc(x_606); +x_653 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__7___boxed), 15, 4); +lean_closure_set(x_653, 0, x_606); +lean_closure_set(x_653, 1, x_649); +lean_closure_set(x_653, 2, x_607); +lean_closure_set(x_653, 3, x_650); +x_654 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__46; +x_655 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg___boxed), 12, 2); +lean_closure_set(x_655, 0, x_654); +lean_closure_set(x_655, 1, x_653); +x_656 = l_Lean_Omega_LinearCombo_add(x_606, x_649); +x_657 = lean_ctor_get(x_651, 1); +lean_inc(x_657); +lean_dec(x_651); +x_658 = lean_array_get_size(x_657); +x_659 = lean_unsigned_to_nat(0u); +x_660 = lean_nat_dec_lt(x_659, x_658); +if (x_660 == 0) +{ +lean_object* x_661; lean_object* x_662; +lean_dec(x_658); +lean_dec(x_657); +if (lean_is_scalar(x_652)) { + x_661 = lean_alloc_ctor(0, 2, 0); } else { - x_659 = x_650; + x_661 = x_652; } -lean_ctor_set(x_659, 0, x_653); -lean_ctor_set(x_659, 1, x_606); -x_660 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_660, 0, x_654); -lean_ctor_set(x_660, 1, x_659); -lean_ctor_set(x_607, 0, x_660); -return x_607; +lean_ctor_set(x_661, 0, x_655); +lean_ctor_set(x_661, 1, x_608); +x_662 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_662, 0, x_656); +lean_ctor_set(x_662, 1, x_661); +lean_ctor_set(x_609, 0, x_662); +return x_609; } else { -uint8_t x_661; -x_661 = lean_nat_dec_le(x_656, x_656); -if (x_661 == 0) +uint8_t x_663; +x_663 = lean_nat_dec_le(x_658, x_658); +if (x_663 == 0) { -lean_object* x_662; lean_object* x_663; -lean_dec(x_656); -lean_dec(x_655); -if (lean_is_scalar(x_650)) { - x_662 = lean_alloc_ctor(0, 2, 0); +lean_object* x_664; lean_object* x_665; +lean_dec(x_658); +lean_dec(x_657); +if (lean_is_scalar(x_652)) { + x_664 = lean_alloc_ctor(0, 2, 0); } else { - x_662 = x_650; -} -lean_ctor_set(x_662, 0, x_653); -lean_ctor_set(x_662, 1, x_606); -x_663 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_663, 0, x_654); -lean_ctor_set(x_663, 1, x_662); -lean_ctor_set(x_607, 0, x_663); -return x_607; -} -else -{ -size_t x_664; size_t x_665; lean_object* x_666; lean_object* x_667; lean_object* x_668; -x_664 = 0; -x_665 = lean_usize_of_nat(x_656); -lean_dec(x_656); -x_666 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Omega_asLinearComboImpl___spec__3(x_655, x_664, x_665, x_606); -lean_dec(x_655); -if (lean_is_scalar(x_650)) { - x_667 = lean_alloc_ctor(0, 2, 0); + x_664 = x_652; +} +lean_ctor_set(x_664, 0, x_655); +lean_ctor_set(x_664, 1, x_608); +x_665 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_665, 0, x_656); +lean_ctor_set(x_665, 1, x_664); +lean_ctor_set(x_609, 0, x_665); +return x_609; +} +else +{ +size_t x_666; size_t x_667; lean_object* x_668; lean_object* x_669; lean_object* x_670; +x_666 = 0; +x_667 = lean_usize_of_nat(x_658); +lean_dec(x_658); +x_668 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Omega_asLinearComboImpl___spec__3(x_657, x_666, x_667, x_608); +lean_dec(x_657); +if (lean_is_scalar(x_652)) { + x_669 = lean_alloc_ctor(0, 2, 0); } else { - x_667 = x_650; + x_669 = x_652; } -lean_ctor_set(x_667, 0, x_653); -lean_ctor_set(x_667, 1, x_666); -x_668 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_668, 0, x_654); -lean_ctor_set(x_668, 1, x_667); -lean_ctor_set(x_607, 0, x_668); -return x_607; +lean_ctor_set(x_669, 0, x_655); +lean_ctor_set(x_669, 1, x_668); +x_670 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_670, 0, x_656); +lean_ctor_set(x_670, 1, x_669); +lean_ctor_set(x_609, 0, x_670); +return x_609; } } } } else { -lean_object* x_669; lean_object* x_670; lean_object* x_671; lean_object* x_672; lean_object* x_673; lean_object* x_674; lean_object* x_675; lean_object* x_676; lean_object* x_677; lean_object* x_678; lean_object* x_679; lean_object* x_680; lean_object* x_681; uint8_t x_682; -x_669 = lean_ctor_get(x_607, 1); -lean_inc(x_669); -lean_dec(x_607); -x_670 = lean_ctor_get(x_608, 0); -lean_inc(x_670); -if (lean_is_exclusive(x_608)) { - lean_ctor_release(x_608, 0); - lean_ctor_release(x_608, 1); - x_671 = x_608; -} else { - lean_dec_ref(x_608); - x_671 = lean_box(0); -} -x_672 = lean_ctor_get(x_609, 0); +lean_object* x_671; lean_object* x_672; lean_object* x_673; lean_object* x_674; lean_object* x_675; lean_object* x_676; lean_object* x_677; lean_object* x_678; lean_object* x_679; lean_object* x_680; lean_object* x_681; lean_object* x_682; lean_object* x_683; uint8_t x_684; +x_671 = lean_ctor_get(x_609, 1); +lean_inc(x_671); +lean_dec(x_609); +x_672 = lean_ctor_get(x_610, 0); lean_inc(x_672); -x_673 = lean_ctor_get(x_609, 1); -lean_inc(x_673); -if (lean_is_exclusive(x_609)) { - lean_ctor_release(x_609, 0); - lean_ctor_release(x_609, 1); - x_674 = x_609; +if (lean_is_exclusive(x_610)) { + lean_ctor_release(x_610, 0); + lean_ctor_release(x_610, 1); + x_673 = x_610; } else { - lean_dec_ref(x_609); - x_674 = lean_box(0); + lean_dec_ref(x_610); + x_673 = lean_box(0); +} +x_674 = lean_ctor_get(x_611, 0); +lean_inc(x_674); +x_675 = lean_ctor_get(x_611, 1); +lean_inc(x_675); +if (lean_is_exclusive(x_611)) { + lean_ctor_release(x_611, 0); + lean_ctor_release(x_611, 1); + x_676 = x_611; +} else { + lean_dec_ref(x_611); + x_676 = lean_box(0); } -lean_inc(x_670); -lean_inc(x_604); -x_675 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__7___boxed), 13, 4); -lean_closure_set(x_675, 0, x_604); -lean_closure_set(x_675, 1, x_670); -lean_closure_set(x_675, 2, x_605); -lean_closure_set(x_675, 3, x_672); -x_676 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__46; -x_677 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg), 10, 2); -lean_closure_set(x_677, 0, x_676); -lean_closure_set(x_677, 1, x_675); -x_678 = l_Lean_Omega_LinearCombo_add(x_604, x_670); -x_679 = lean_ctor_get(x_673, 1); -lean_inc(x_679); -lean_dec(x_673); -x_680 = lean_array_get_size(x_679); -x_681 = lean_unsigned_to_nat(0u); -x_682 = lean_nat_dec_lt(x_681, x_680); -if (x_682 == 0) -{ -lean_object* x_683; lean_object* x_684; lean_object* x_685; -lean_dec(x_680); -lean_dec(x_679); -if (lean_is_scalar(x_674)) { - x_683 = lean_alloc_ctor(0, 2, 0); +lean_inc(x_672); +lean_inc(x_606); +x_677 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__7___boxed), 15, 4); +lean_closure_set(x_677, 0, x_606); +lean_closure_set(x_677, 1, x_672); +lean_closure_set(x_677, 2, x_607); +lean_closure_set(x_677, 3, x_674); +x_678 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__46; +x_679 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg___boxed), 12, 2); +lean_closure_set(x_679, 0, x_678); +lean_closure_set(x_679, 1, x_677); +x_680 = l_Lean_Omega_LinearCombo_add(x_606, x_672); +x_681 = lean_ctor_get(x_675, 1); +lean_inc(x_681); +lean_dec(x_675); +x_682 = lean_array_get_size(x_681); +x_683 = lean_unsigned_to_nat(0u); +x_684 = lean_nat_dec_lt(x_683, x_682); +if (x_684 == 0) +{ +lean_object* x_685; lean_object* x_686; lean_object* x_687; +lean_dec(x_682); +lean_dec(x_681); +if (lean_is_scalar(x_676)) { + x_685 = lean_alloc_ctor(0, 2, 0); } else { - x_683 = x_674; + x_685 = x_676; } -lean_ctor_set(x_683, 0, x_677); -lean_ctor_set(x_683, 1, x_606); -if (lean_is_scalar(x_671)) { - x_684 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_685, 0, x_679); +lean_ctor_set(x_685, 1, x_608); +if (lean_is_scalar(x_673)) { + x_686 = lean_alloc_ctor(0, 2, 0); } else { - x_684 = x_671; + x_686 = x_673; } -lean_ctor_set(x_684, 0, x_678); -lean_ctor_set(x_684, 1, x_683); -x_685 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_685, 0, x_684); -lean_ctor_set(x_685, 1, x_669); -return x_685; +lean_ctor_set(x_686, 0, x_680); +lean_ctor_set(x_686, 1, x_685); +x_687 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_687, 0, x_686); +lean_ctor_set(x_687, 1, x_671); +return x_687; } else { -uint8_t x_686; -x_686 = lean_nat_dec_le(x_680, x_680); -if (x_686 == 0) +uint8_t x_688; +x_688 = lean_nat_dec_le(x_682, x_682); +if (x_688 == 0) { -lean_object* x_687; lean_object* x_688; lean_object* x_689; -lean_dec(x_680); -lean_dec(x_679); -if (lean_is_scalar(x_674)) { - x_687 = lean_alloc_ctor(0, 2, 0); +lean_object* x_689; lean_object* x_690; lean_object* x_691; +lean_dec(x_682); +lean_dec(x_681); +if (lean_is_scalar(x_676)) { + x_689 = lean_alloc_ctor(0, 2, 0); } else { - x_687 = x_674; + x_689 = x_676; } -lean_ctor_set(x_687, 0, x_677); -lean_ctor_set(x_687, 1, x_606); -if (lean_is_scalar(x_671)) { - x_688 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_689, 0, x_679); +lean_ctor_set(x_689, 1, x_608); +if (lean_is_scalar(x_673)) { + x_690 = lean_alloc_ctor(0, 2, 0); } else { - x_688 = x_671; + x_690 = x_673; } -lean_ctor_set(x_688, 0, x_678); -lean_ctor_set(x_688, 1, x_687); -x_689 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_689, 0, x_688); -lean_ctor_set(x_689, 1, x_669); -return x_689; +lean_ctor_set(x_690, 0, x_680); +lean_ctor_set(x_690, 1, x_689); +x_691 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_691, 0, x_690); +lean_ctor_set(x_691, 1, x_671); +return x_691; } else { -size_t x_690; size_t x_691; lean_object* x_692; lean_object* x_693; lean_object* x_694; lean_object* x_695; -x_690 = 0; -x_691 = lean_usize_of_nat(x_680); -lean_dec(x_680); -x_692 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Omega_asLinearComboImpl___spec__3(x_679, x_690, x_691, x_606); -lean_dec(x_679); -if (lean_is_scalar(x_674)) { - x_693 = lean_alloc_ctor(0, 2, 0); +size_t x_692; size_t x_693; lean_object* x_694; lean_object* x_695; lean_object* x_696; lean_object* x_697; +x_692 = 0; +x_693 = lean_usize_of_nat(x_682); +lean_dec(x_682); +x_694 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Omega_asLinearComboImpl___spec__3(x_681, x_692, x_693, x_608); +lean_dec(x_681); +if (lean_is_scalar(x_676)) { + x_695 = lean_alloc_ctor(0, 2, 0); } else { - x_693 = x_674; + x_695 = x_676; } -lean_ctor_set(x_693, 0, x_677); -lean_ctor_set(x_693, 1, x_692); -if (lean_is_scalar(x_671)) { - x_694 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_695, 0, x_679); +lean_ctor_set(x_695, 1, x_694); +if (lean_is_scalar(x_673)) { + x_696 = lean_alloc_ctor(0, 2, 0); } else { - x_694 = x_671; + x_696 = x_673; } -lean_ctor_set(x_694, 0, x_678); -lean_ctor_set(x_694, 1, x_693); -x_695 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_695, 0, x_694); -lean_ctor_set(x_695, 1, x_669); -return x_695; +lean_ctor_set(x_696, 0, x_680); +lean_ctor_set(x_696, 1, x_695); +x_697 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_697, 0, x_696); +lean_ctor_set(x_697, 1, x_671); +return x_697; } } } } else { -uint8_t x_696; +uint8_t x_698; +lean_dec(x_608); +lean_dec(x_607); lean_dec(x_606); -lean_dec(x_605); -lean_dec(x_604); -x_696 = !lean_is_exclusive(x_607); -if (x_696 == 0) +x_698 = !lean_is_exclusive(x_609); +if (x_698 == 0) { -return x_607; +return x_609; } else { -lean_object* x_697; lean_object* x_698; lean_object* x_699; -x_697 = lean_ctor_get(x_607, 0); -x_698 = lean_ctor_get(x_607, 1); -lean_inc(x_698); -lean_inc(x_697); -lean_dec(x_607); -x_699 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_699, 0, x_697); -lean_ctor_set(x_699, 1, x_698); -return x_699; +lean_object* x_699; lean_object* x_700; lean_object* x_701; +x_699 = lean_ctor_get(x_609, 0); +x_700 = lean_ctor_get(x_609, 1); +lean_inc(x_700); +lean_inc(x_699); +lean_dec(x_609); +x_701 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_701, 0, x_699); +lean_ctor_set(x_701, 1, x_700); +return x_701; } } } else { -uint8_t x_700; -lean_dec(x_599); +uint8_t x_702; +lean_dec(x_601); +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_700 = !lean_is_exclusive(x_600); -if (x_700 == 0) +x_702 = !lean_is_exclusive(x_602); +if (x_702 == 0) { -return x_600; +return x_602; } else { -lean_object* x_701; lean_object* x_702; lean_object* x_703; -x_701 = lean_ctor_get(x_600, 0); -x_702 = lean_ctor_get(x_600, 1); -lean_inc(x_702); -lean_inc(x_701); -lean_dec(x_600); -x_703 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_703, 0, x_701); -lean_ctor_set(x_703, 1, x_702); -return x_703; +lean_object* x_703; lean_object* x_704; lean_object* x_705; +x_703 = lean_ctor_get(x_602, 0); +x_704 = lean_ctor_get(x_602, 1); +lean_inc(x_704); +lean_inc(x_703); +lean_dec(x_602); +x_705 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_705, 0, x_703); +lean_ctor_set(x_705, 1, x_704); +return x_705; } } } @@ -11861,177 +11968,180 @@ return x_703; } else { -lean_object* x_704; +lean_object* x_706; +lean_dec(x_18); +lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -x_704 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_704; +x_706 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_706; } } else { -lean_object* x_705; +lean_object* x_707; +lean_dec(x_17); +lean_dec(x_16); lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -x_705 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_705; +x_707 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_707; } } else { -lean_object* x_706; -lean_dec(x_14); -lean_dec(x_13); -x_706 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_706; +lean_object* x_708; +lean_dec(x_16); +lean_dec(x_15); +x_708 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_708; } } else { -lean_object* x_707; lean_object* x_708; +lean_object* x_709; lean_object* x_710; lean_inc(x_1); -x_707 = l_Lean_Expr_fvarId_x21(x_1); -lean_inc(x_6); -x_708 = l_Lean_FVarId_getValue_x3f(x_707, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_708) == 0) +x_709 = l_Lean_Expr_fvarId_x21(x_1); +lean_inc(x_8); +x_710 = l_Lean_FVarId_getValue_x3f(x_709, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_710) == 0) { -lean_object* x_709; -x_709 = lean_ctor_get(x_708, 0); -lean_inc(x_709); -if (lean_obj_tag(x_709) == 0) -{ -lean_object* x_710; lean_object* x_711; -x_710 = lean_ctor_get(x_708, 1); -lean_inc(x_710); -lean_dec(x_708); -x_711 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_710); -return x_711; +lean_object* x_711; +x_711 = lean_ctor_get(x_710, 0); +lean_inc(x_711); +if (lean_obj_tag(x_711) == 0) +{ +lean_object* x_712; lean_object* x_713; +x_712 = lean_ctor_get(x_710, 1); +lean_inc(x_712); +lean_dec(x_710); +x_713 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_712); +return x_713; } else { -lean_object* x_712; lean_object* x_713; lean_object* x_714; -x_712 = lean_ctor_get(x_708, 1); -lean_inc(x_712); -lean_dec(x_708); -x_713 = lean_ctor_get(x_709, 0); -lean_inc(x_713); -lean_dec(x_709); +lean_object* x_714; lean_object* x_715; lean_object* x_716; +x_714 = lean_ctor_get(x_710, 1); +lean_inc(x_714); +lean_dec(x_710); +x_715 = lean_ctor_get(x_711, 0); +lean_inc(x_715); +lean_dec(x_711); +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_714 = l_Lean_Elab_Tactic_Omega_mkEqReflWithExpectedType(x_1, x_713, x_6, x_7, x_8, x_9, x_712); -if (lean_obj_tag(x_714) == 0) -{ -lean_object* x_715; lean_object* x_716; lean_object* x_717; -x_715 = lean_ctor_get(x_714, 0); -lean_inc(x_715); -x_716 = lean_ctor_get(x_714, 1); -lean_inc(x_716); -lean_dec(x_714); -x_717 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_715, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_716); -return x_717; +x_716 = l_Lean_Elab_Tactic_Omega_mkEqReflWithExpectedType(x_1, x_715, x_8, x_9, x_10, x_11, x_714); +if (lean_obj_tag(x_716) == 0) +{ +lean_object* x_717; lean_object* x_718; lean_object* x_719; +x_717 = lean_ctor_get(x_716, 0); +lean_inc(x_717); +x_718 = lean_ctor_get(x_716, 1); +lean_inc(x_718); +lean_dec(x_716); +x_719 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_717, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_718); +return x_719; } else { -uint8_t x_718; +uint8_t x_720; +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_1); -x_718 = !lean_is_exclusive(x_714); -if (x_718 == 0) +x_720 = !lean_is_exclusive(x_716); +if (x_720 == 0) { -return x_714; +return x_716; } else { -lean_object* x_719; lean_object* x_720; lean_object* x_721; -x_719 = lean_ctor_get(x_714, 0); -x_720 = lean_ctor_get(x_714, 1); -lean_inc(x_720); -lean_inc(x_719); -lean_dec(x_714); -x_721 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_721, 0, x_719); -lean_ctor_set(x_721, 1, x_720); -return x_721; +lean_object* x_721; lean_object* x_722; lean_object* x_723; +x_721 = lean_ctor_get(x_716, 0); +x_722 = lean_ctor_get(x_716, 1); +lean_inc(x_722); +lean_inc(x_721); +lean_dec(x_716); +x_723 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_723, 0, x_721); +lean_ctor_set(x_723, 1, x_722); +return x_723; } } } } else { -uint8_t x_722; +uint8_t x_724; +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_1); -x_722 = !lean_is_exclusive(x_708); -if (x_722 == 0) +x_724 = !lean_is_exclusive(x_710); +if (x_724 == 0) { -return x_708; +return x_710; } else { -lean_object* x_723; lean_object* x_724; lean_object* x_725; -x_723 = lean_ctor_get(x_708, 0); -x_724 = lean_ctor_get(x_708, 1); -lean_inc(x_724); -lean_inc(x_723); -lean_dec(x_708); -x_725 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_725, 0, x_723); -lean_ctor_set(x_725, 1, x_724); -return x_725; +lean_object* x_725; lean_object* x_726; lean_object* x_727; +x_725 = lean_ctor_get(x_710, 0); +x_726 = lean_ctor_get(x_710, 1); +lean_inc(x_726); +lean_inc(x_725); +lean_dec(x_710); +x_727 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_727, 0, x_725); +lean_ctor_set(x_727, 1, x_726); +return x_727; } } } } else { -lean_object* x_726; lean_object* x_727; lean_object* x_728; lean_object* x_729; lean_object* x_730; lean_object* x_731; lean_object* x_732; lean_object* x_733; +lean_object* x_728; lean_object* x_729; lean_object* x_730; lean_object* x_731; lean_object* x_732; lean_object* x_733; lean_object* x_734; lean_object* x_735; +lean_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_726 = lean_ctor_get(x_11, 0); -lean_inc(x_726); -lean_dec(x_11); -x_727 = lean_box(0); -x_728 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_728, 0, x_726); -lean_ctor_set(x_728, 1, x_727); +x_728 = lean_ctor_get(x_13, 0); lean_inc(x_728); -x_729 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_mkEvalRflProof___boxed), 10, 2); -lean_closure_set(x_729, 0, x_1); -lean_closure_set(x_729, 1, x_728); -x_730 = l_Lean_Elab_Tactic_Omega_MetaProblem_problem___default___closed__2; -x_731 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_731, 0, x_729); -lean_ctor_set(x_731, 1, x_730); -x_732 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_732, 0, x_728); -lean_ctor_set(x_732, 1, x_731); +lean_dec(x_13); +x_729 = lean_box(0); +x_730 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_730, 0, x_728); +lean_ctor_set(x_730, 1, x_729); +lean_inc(x_730); +x_731 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_mkEvalRflProof___boxed), 12, 2); +lean_closure_set(x_731, 0, x_1); +lean_closure_set(x_731, 1, x_730); +x_732 = l_Lean_Elab_Tactic_Omega_MetaProblem_problem___default___closed__2; x_733 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_733, 0, x_732); -lean_ctor_set(x_733, 1, x_10); -return x_733; +lean_ctor_set(x_733, 0, x_731); +lean_ctor_set(x_733, 1, x_732); +x_734 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_734, 0, x_730); +lean_ctor_set(x_734, 1, x_733); +x_735 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_735, 0, x_734); +lean_ctor_set(x_735, 1, x_12); +return x_735; } } } @@ -12052,50 +12162,50 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl(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) { _start: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_10 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__2; -x_11 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_10, x_2, x_3, x_4, 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); -lean_dec(x_12); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_11, 1); +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_12 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__2; +x_13 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_12, x_2, x_3, x_4, 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); -lean_dec(x_11); -x_15 = lean_box(0); -x_16 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8(x_1, x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_14); -return x_16; +x_15 = lean_unbox(x_14); +lean_dec(x_14); +if (x_15 == 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_dec(x_13); +x_17 = lean_box(0); +x_18 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8(x_1, x_17, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_16); +return x_18; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_17 = lean_ctor_get(x_11, 1); -lean_inc(x_17); -lean_dec(x_11); +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_19 = lean_ctor_get(x_13, 1); +lean_inc(x_19); +lean_dec(x_13); lean_inc(x_1); -x_18 = l_Lean_MessageData_ofExpr(x_1); -x_19 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___closed__2; -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_Tactic_Omega_asLinearComboImpl_rewrite___closed__9; +x_20 = l_Lean_MessageData_ofExpr(x_1); +x_21 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___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 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_10, x_22, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_17); -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 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8(x_1, x_24, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_25); -return x_26; +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_20); +x_23 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__9; +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_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_12, x_24, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_19); +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_Elab_Tactic_Omega_asLinearComboImpl___lambda__8(x_1, x_26, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_27); +return x_28; } } } @@ -12611,328 +12721,330 @@ x_3 = l_Lean_Expr_const___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_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_Omega_asLinearComboImpl_handleNatCast(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_object* x_12, lean_object* x_13) { _start: { if (lean_obj_tag(x_3) == 1) { -lean_object* x_12; lean_object* x_13; -x_12 = lean_ctor_get(x_3, 0); -lean_inc(x_12); +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_3, 0); +lean_inc(x_14); lean_dec(x_3); -lean_inc(x_7); -x_13 = l_Lean_FVarId_getValue_x3f(x_12, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_13) == 0) +lean_inc(x_9); +x_15 = l_Lean_FVarId_getValue_x3f(x_14, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_15) == 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_16; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_15; lean_object* x_16; +lean_object* x_17; lean_object* x_18; lean_dec(x_2); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); -x_16 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_15); -return x_16; +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_17); +return x_18; } 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; -x_17 = lean_ctor_get(x_13, 1); -lean_inc(x_17); -lean_dec(x_13); -x_18 = lean_ctor_get(x_14, 0); -lean_inc(x_18); -lean_dec(x_14); -x_19 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__54; -x_20 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__11; -x_21 = l_Lean_mkApp3(x_19, x_20, x_2, x_18); +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_19 = lean_ctor_get(x_15, 1); +lean_inc(x_19); +lean_dec(x_15); +x_20 = lean_ctor_get(x_16, 0); +lean_inc(x_20); +lean_dec(x_16); +x_21 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__54; +x_22 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__11; +x_23 = l_Lean_mkApp3(x_21, x_22, x_2, x_20); +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); -x_22 = l_Lean_Elab_Tactic_Omega_mkEqReflWithExpectedType(x_1, x_21, x_7, x_8, x_9, x_10, x_17); -if (lean_obj_tag(x_22) == 0) +x_24 = l_Lean_Elab_Tactic_Omega_mkEqReflWithExpectedType(x_1, x_23, x_9, x_10, x_11, x_12, x_19); +if (lean_obj_tag(x_24) == 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); -x_25 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_23, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_24); -return x_25; +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +lean_dec(x_24); +x_27 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_25, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_26); +return x_27; } else { -uint8_t x_26; +uint8_t x_28; +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_1); -x_26 = !lean_is_exclusive(x_22); -if (x_26 == 0) +x_28 = !lean_is_exclusive(x_24); +if (x_28 == 0) { -return x_22; +return x_24; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_22, 0); -x_28 = lean_ctor_get(x_22, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_22); -x_29 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_28); -return x_29; +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_24, 0); +x_30 = lean_ctor_get(x_24, 1); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_24); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +return x_31; } } } } else { -uint8_t x_30; +uint8_t x_32; +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_2); lean_dec(x_1); -x_30 = !lean_is_exclusive(x_13); -if (x_30 == 0) +x_32 = !lean_is_exclusive(x_15); +if (x_32 == 0) { -return x_13; +return x_15; } 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_dec(x_13); -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_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_15, 0); +x_34 = lean_ctor_get(x_15, 1); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_15); +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_34; lean_object* x_35; -x_34 = l_Lean_Expr_getAppFnArgs(x_3); -x_35 = lean_ctor_get(x_34, 0); -lean_inc(x_35); -if (lean_obj_tag(x_35) == 1) -{ -lean_object* x_36; -x_36 = lean_ctor_get(x_35, 0); -lean_inc(x_36); -if (lean_obj_tag(x_36) == 1) -{ -lean_object* x_37; +lean_object* x_36; lean_object* x_37; +x_36 = l_Lean_Expr_getAppFnArgs(x_3); x_37 = lean_ctor_get(x_36, 0); lean_inc(x_37); -if (lean_obj_tag(x_37) == 0) +if (lean_obj_tag(x_37) == 1) { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; -x_38 = lean_ctor_get(x_34, 1); +lean_object* x_38; +x_38 = lean_ctor_get(x_37, 0); lean_inc(x_38); -lean_dec(x_34); -x_39 = lean_ctor_get(x_35, 1); +if (lean_obj_tag(x_38) == 1) +{ +lean_object* x_39; +x_39 = lean_ctor_get(x_38, 0); lean_inc(x_39); -lean_dec(x_35); +if (lean_obj_tag(x_39) == 0) +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; x_40 = lean_ctor_get(x_36, 1); lean_inc(x_40); lean_dec(x_36); -x_41 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__1; -x_42 = lean_string_dec_eq(x_40, x_41); -if (x_42 == 0) -{ -lean_object* x_43; uint8_t x_44; -x_43 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__2; -x_44 = lean_string_dec_eq(x_40, x_43); +x_41 = lean_ctor_get(x_37, 1); +lean_inc(x_41); +lean_dec(x_37); +x_42 = lean_ctor_get(x_38, 1); +lean_inc(x_42); +lean_dec(x_38); +x_43 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__1; +x_44 = lean_string_dec_eq(x_42, x_43); if (x_44 == 0) { lean_object* x_45; uint8_t x_46; -x_45 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__2; -x_46 = lean_string_dec_eq(x_40, x_45); +x_45 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__2; +x_46 = lean_string_dec_eq(x_42, x_45); if (x_46 == 0) { lean_object* x_47; uint8_t x_48; -x_47 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__4; -x_48 = lean_string_dec_eq(x_40, x_47); +x_47 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__2; +x_48 = lean_string_dec_eq(x_42, x_47); if (x_48 == 0) { lean_object* x_49; uint8_t x_50; -x_49 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__1; -x_50 = lean_string_dec_eq(x_40, x_49); +x_49 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__4; +x_50 = lean_string_dec_eq(x_42, x_49); if (x_50 == 0) { lean_object* x_51; uint8_t x_52; -x_51 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__3; -x_52 = lean_string_dec_eq(x_40, x_51); +x_51 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__1; +x_52 = lean_string_dec_eq(x_42, x_51); if (x_52 == 0) { lean_object* x_53; uint8_t x_54; -x_53 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__1; -x_54 = lean_string_dec_eq(x_40, x_53); +x_53 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__3; +x_54 = lean_string_dec_eq(x_42, x_53); if (x_54 == 0) { lean_object* x_55; uint8_t x_56; -x_55 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__7; -x_56 = lean_string_dec_eq(x_40, x_55); +x_55 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__1; +x_56 = lean_string_dec_eq(x_42, x_55); if (x_56 == 0) { lean_object* x_57; uint8_t x_58; -x_57 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__8; -x_58 = lean_string_dec_eq(x_40, x_57); +x_57 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__7; +x_58 = lean_string_dec_eq(x_42, x_57); if (x_58 == 0) { lean_object* x_59; uint8_t x_60; -x_59 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__5; -x_60 = lean_string_dec_eq(x_40, x_59); +x_59 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__8; +x_60 = lean_string_dec_eq(x_42, x_59); if (x_60 == 0) { lean_object* x_61; uint8_t x_62; -x_61 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__6; -x_62 = lean_string_dec_eq(x_40, x_61); +x_61 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__5; +x_62 = lean_string_dec_eq(x_42, x_61); if (x_62 == 0) { lean_object* x_63; uint8_t x_64; -x_63 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__2; -x_64 = lean_string_dec_eq(x_40, x_63); +x_63 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__6; +x_64 = lean_string_dec_eq(x_42, x_63); if (x_64 == 0) { lean_object* x_65; uint8_t x_66; -x_65 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__3; -x_66 = lean_string_dec_eq(x_40, x_65); +x_65 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__2; +x_66 = lean_string_dec_eq(x_42, x_65); if (x_66 == 0) { lean_object* x_67; uint8_t x_68; -x_67 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__9; -x_68 = lean_string_dec_eq(x_40, x_67); +x_67 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__3; +x_68 = lean_string_dec_eq(x_42, x_67); if (x_68 == 0) { lean_object* x_69; uint8_t x_70; -x_69 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__4; -x_70 = lean_string_dec_eq(x_40, x_69); -lean_dec(x_40); +x_69 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__9; +x_70 = lean_string_dec_eq(x_42, x_69); if (x_70 == 0) { -lean_object* x_71; -lean_dec(x_39); -lean_dec(x_38); +lean_object* x_71; uint8_t x_72; +x_71 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__4; +x_72 = lean_string_dec_eq(x_42, x_71); +lean_dec(x_42); +if (x_72 == 0) +{ +lean_object* x_73; +lean_dec(x_41); +lean_dec(x_40); lean_dec(x_2); -x_71 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_71; +x_73 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_73; } else { -lean_object* x_72; uint8_t x_73; -x_72 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__5; -x_73 = lean_string_dec_eq(x_39, x_72); -lean_dec(x_39); -if (x_73 == 0) +lean_object* x_74; uint8_t x_75; +x_74 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__5; +x_75 = lean_string_dec_eq(x_41, x_74); +lean_dec(x_41); +if (x_75 == 0) { -lean_object* x_74; -lean_dec(x_38); +lean_object* x_76; +lean_dec(x_40); lean_dec(x_2); -x_74 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_74; +x_76 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_76; } else { -lean_object* x_75; lean_object* x_76; uint8_t x_77; -x_75 = lean_array_get_size(x_38); -x_76 = lean_unsigned_to_nat(2u); -x_77 = lean_nat_dec_eq(x_75, x_76); -lean_dec(x_75); -if (x_77 == 0) +lean_object* x_77; lean_object* x_78; uint8_t x_79; +x_77 = lean_array_get_size(x_40); +x_78 = lean_unsigned_to_nat(2u); +x_79 = lean_nat_dec_eq(x_77, x_78); +lean_dec(x_77); +if (x_79 == 0) { -lean_object* x_78; -lean_dec(x_38); +lean_object* x_80; +lean_dec(x_40); lean_dec(x_2); -x_78 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_78; +x_80 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_80; } else { -lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; -x_79 = lean_unsigned_to_nat(0u); -x_80 = lean_array_fget(x_38, x_79); -x_81 = lean_unsigned_to_nat(1u); -x_82 = lean_array_fget(x_38, x_81); -lean_dec(x_38); -x_83 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleFinVal(x_1, x_2, x_80, x_82, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_83; +lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_81 = lean_unsigned_to_nat(0u); +x_82 = lean_array_fget(x_40, x_81); +x_83 = lean_unsigned_to_nat(1u); +x_84 = lean_array_fget(x_40, x_83); +lean_dec(x_40); +x_85 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleFinVal(x_1, x_2, x_82, x_84, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_85; } } } } else { -lean_object* x_84; uint8_t x_85; -lean_dec(x_40); +lean_object* x_86; uint8_t x_87; +lean_dec(x_42); lean_dec(x_2); -x_84 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__6; -x_85 = lean_string_dec_eq(x_39, x_84); -lean_dec(x_39); -if (x_85 == 0) +x_86 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__6; +x_87 = lean_string_dec_eq(x_41, x_86); +lean_dec(x_41); +if (x_87 == 0) { -lean_object* x_86; -lean_dec(x_38); -x_86 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_86; +lean_object* x_88; +lean_dec(x_40); +x_88 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_88; } else { -lean_object* x_87; lean_object* x_88; uint8_t x_89; -x_87 = lean_array_get_size(x_38); -x_88 = lean_unsigned_to_nat(1u); -x_89 = lean_nat_dec_eq(x_87, x_88); -lean_dec(x_87); -if (x_89 == 0) -{ -lean_object* x_90; -lean_dec(x_38); -x_90 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_90; -} -else -{ -uint8_t x_91; -x_91 = lean_ctor_get_uint8(x_6, 2); -if (x_91 == 0) +lean_object* x_89; lean_object* x_90; uint8_t x_91; +x_89 = lean_array_get_size(x_40); +x_90 = lean_unsigned_to_nat(1u); +x_91 = lean_nat_dec_eq(x_89, x_90); +lean_dec(x_89); +if (x_91 == 0) { lean_object* x_92; -lean_dec(x_38); -x_92 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_40); +x_92 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); return x_92; } else { -lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; -x_93 = lean_unsigned_to_nat(0u); -x_94 = lean_array_fget(x_38, x_93); -lean_dec(x_38); -x_95 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__9; -x_96 = l_Lean_Expr_app___override(x_95, x_94); -x_97 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_96, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_97; +uint8_t x_93; +x_93 = lean_ctor_get_uint8(x_6, 2); +if (x_93 == 0) +{ +lean_object* x_94; +lean_dec(x_40); +x_94 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_94; +} +else +{ +lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_95 = lean_unsigned_to_nat(0u); +x_96 = lean_array_fget(x_40, x_95); +lean_dec(x_40); +x_97 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__9; +x_98 = l_Lean_Expr_app___override(x_97, x_96); +x_99 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_98, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_99; } } } @@ -12940,425 +13052,425 @@ return x_97; } else { -lean_object* x_98; uint8_t x_99; -lean_dec(x_40); +lean_object* x_100; uint8_t x_101; +lean_dec(x_42); lean_dec(x_2); -x_98 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__10; -x_99 = lean_string_dec_eq(x_39, x_98); -lean_dec(x_39); -if (x_99 == 0) +x_100 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__10; +x_101 = lean_string_dec_eq(x_41, x_100); +lean_dec(x_41); +if (x_101 == 0) { -lean_object* x_100; -lean_dec(x_38); -x_100 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_100; +lean_object* x_102; +lean_dec(x_40); +x_102 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_102; } else { -lean_object* x_101; lean_object* x_102; uint8_t x_103; -x_101 = lean_array_get_size(x_38); -x_102 = lean_unsigned_to_nat(6u); -x_103 = lean_nat_dec_eq(x_101, x_102); -lean_dec(x_101); -if (x_103 == 0) +lean_object* x_103; lean_object* x_104; uint8_t x_105; +x_103 = lean_array_get_size(x_40); +x_104 = lean_unsigned_to_nat(6u); +x_105 = lean_nat_dec_eq(x_103, x_104); +lean_dec(x_103); +if (x_105 == 0) { -lean_object* x_104; -lean_dec(x_38); -x_104 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_104; +lean_object* x_106; +lean_dec(x_40); +x_106 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_106; } else { -lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; -x_105 = lean_unsigned_to_nat(4u); -x_106 = lean_array_fget(x_38, x_105); -x_107 = lean_unsigned_to_nat(5u); -x_108 = lean_array_fget(x_38, x_107); -lean_dec(x_38); -x_109 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__13; -x_110 = l_Lean_mkAppB(x_109, x_106, x_108); -x_111 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_110, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_111; +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_unsigned_to_nat(4u); +x_108 = lean_array_fget(x_40, x_107); +x_109 = lean_unsigned_to_nat(5u); +x_110 = lean_array_fget(x_40, x_109); +lean_dec(x_40); +x_111 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__13; +x_112 = l_Lean_mkAppB(x_111, x_108, x_110); +x_113 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_112, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_113; } } } } else { -lean_object* x_112; uint8_t x_113; -lean_dec(x_40); +lean_object* x_114; uint8_t x_115; +lean_dec(x_42); lean_dec(x_2); -x_112 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__14; -x_113 = lean_string_dec_eq(x_39, x_112); -lean_dec(x_39); -if (x_113 == 0) +x_114 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__14; +x_115 = lean_string_dec_eq(x_41, x_114); +lean_dec(x_41); +if (x_115 == 0) { -lean_object* x_114; -lean_dec(x_38); -x_114 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_114; +lean_object* x_116; +lean_dec(x_40); +x_116 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_116; } else { -lean_object* x_115; lean_object* x_116; uint8_t x_117; -x_115 = lean_array_get_size(x_38); -x_116 = lean_unsigned_to_nat(6u); -x_117 = lean_nat_dec_eq(x_115, x_116); -lean_dec(x_115); -if (x_117 == 0) +lean_object* x_117; lean_object* x_118; uint8_t x_119; +x_117 = lean_array_get_size(x_40); +x_118 = lean_unsigned_to_nat(6u); +x_119 = lean_nat_dec_eq(x_117, x_118); +lean_dec(x_117); +if (x_119 == 0) { -lean_object* x_118; -lean_dec(x_38); -x_118 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_118; +lean_object* x_120; +lean_dec(x_40); +x_120 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_120; } else { -lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; -x_119 = lean_unsigned_to_nat(4u); -x_120 = lean_array_fget(x_38, x_119); -x_121 = lean_unsigned_to_nat(5u); -x_122 = lean_array_fget(x_38, x_121); -lean_dec(x_38); -x_123 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__17; -x_124 = l_Lean_mkAppB(x_123, x_120, x_122); -x_125 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_124, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_125; +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_unsigned_to_nat(4u); +x_122 = lean_array_fget(x_40, x_121); +x_123 = lean_unsigned_to_nat(5u); +x_124 = lean_array_fget(x_40, x_123); +lean_dec(x_40); +x_125 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__17; +x_126 = l_Lean_mkAppB(x_125, x_122, x_124); +x_127 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_126, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_127; } } } } else { -lean_object* x_126; uint8_t x_127; -lean_dec(x_40); +lean_object* x_128; uint8_t x_129; +lean_dec(x_42); lean_dec(x_2); -x_126 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__20; -x_127 = lean_string_dec_eq(x_39, x_126); -lean_dec(x_39); -if (x_127 == 0) +x_128 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__20; +x_129 = lean_string_dec_eq(x_41, x_128); +lean_dec(x_41); +if (x_129 == 0) { -lean_object* x_128; -lean_dec(x_38); -x_128 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_128; +lean_object* x_130; +lean_dec(x_40); +x_130 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_130; } else { -lean_object* x_129; lean_object* x_130; uint8_t x_131; -x_129 = lean_array_get_size(x_38); -x_130 = lean_unsigned_to_nat(4u); -x_131 = lean_nat_dec_eq(x_129, x_130); -lean_dec(x_129); -if (x_131 == 0) +lean_object* x_131; lean_object* x_132; uint8_t x_133; +x_131 = lean_array_get_size(x_40); +x_132 = lean_unsigned_to_nat(4u); +x_133 = lean_nat_dec_eq(x_131, x_132); +lean_dec(x_131); +if (x_133 == 0) { -lean_object* x_132; -lean_dec(x_38); -x_132 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_132; +lean_object* x_134; +lean_dec(x_40); +x_134 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_134; } else { -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; -x_133 = lean_unsigned_to_nat(2u); -x_134 = lean_array_fget(x_38, x_133); -x_135 = lean_unsigned_to_nat(3u); -x_136 = lean_array_fget(x_38, x_135); -lean_dec(x_38); -x_137 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__20; -x_138 = l_Lean_mkAppB(x_137, x_134, x_136); -x_139 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_138, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_139; +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_135 = lean_unsigned_to_nat(2u); +x_136 = lean_array_fget(x_40, x_135); +x_137 = lean_unsigned_to_nat(3u); +x_138 = lean_array_fget(x_40, x_137); +lean_dec(x_40); +x_139 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__20; +x_140 = l_Lean_mkAppB(x_139, x_136, x_138); +x_141 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_140, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_141; } } } } else { -lean_object* x_140; uint8_t x_141; -lean_dec(x_40); +lean_object* x_142; uint8_t x_143; +lean_dec(x_42); lean_dec(x_2); -x_140 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__24; -x_141 = lean_string_dec_eq(x_39, x_140); -lean_dec(x_39); -if (x_141 == 0) +x_142 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__24; +x_143 = lean_string_dec_eq(x_41, x_142); +lean_dec(x_41); +if (x_143 == 0) { -lean_object* x_142; -lean_dec(x_38); -x_142 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_142; +lean_object* x_144; +lean_dec(x_40); +x_144 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_144; } else { -lean_object* x_143; lean_object* x_144; uint8_t x_145; -x_143 = lean_array_get_size(x_38); -x_144 = lean_unsigned_to_nat(4u); -x_145 = lean_nat_dec_eq(x_143, x_144); -lean_dec(x_143); -if (x_145 == 0) +lean_object* x_145; lean_object* x_146; uint8_t x_147; +x_145 = lean_array_get_size(x_40); +x_146 = lean_unsigned_to_nat(4u); +x_147 = lean_nat_dec_eq(x_145, x_146); +lean_dec(x_145); +if (x_147 == 0) { -lean_object* x_146; -lean_dec(x_38); -x_146 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_146; +lean_object* x_148; +lean_dec(x_40); +x_148 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_148; } else { -lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; -x_147 = lean_unsigned_to_nat(2u); -x_148 = lean_array_fget(x_38, x_147); -x_149 = lean_unsigned_to_nat(3u); -x_150 = lean_array_fget(x_38, x_149); -lean_dec(x_38); -x_151 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__23; -x_152 = l_Lean_mkAppB(x_151, x_148, x_150); -x_153 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_152, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_153; +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_149 = lean_unsigned_to_nat(2u); +x_150 = lean_array_fget(x_40, x_149); +x_151 = lean_unsigned_to_nat(3u); +x_152 = lean_array_fget(x_40, x_151); +lean_dec(x_40); +x_153 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__23; +x_154 = l_Lean_mkAppB(x_153, x_150, x_152); +x_155 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_154, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_155; } } } } else { -lean_object* x_154; uint8_t x_155; -lean_dec(x_40); -lean_dec(x_2); -x_154 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__9; -x_155 = lean_string_dec_eq(x_39, x_154); -if (x_155 == 0) -{ lean_object* x_156; uint8_t x_157; -x_156 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__10; -x_157 = lean_string_dec_eq(x_39, x_156); -lean_dec(x_39); +lean_dec(x_42); +lean_dec(x_2); +x_156 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__9; +x_157 = lean_string_dec_eq(x_41, x_156); if (x_157 == 0) { -lean_object* x_158; -lean_dec(x_38); -x_158 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_158; +lean_object* x_158; uint8_t x_159; +x_158 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__10; +x_159 = lean_string_dec_eq(x_41, x_158); +lean_dec(x_41); +if (x_159 == 0) +{ +lean_object* x_160; +lean_dec(x_40); +x_160 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_160; } else { -lean_object* x_159; lean_object* x_160; uint8_t x_161; -x_159 = lean_array_get_size(x_38); -x_160 = lean_unsigned_to_nat(3u); -x_161 = lean_nat_dec_eq(x_159, x_160); -lean_dec(x_159); -if (x_161 == 0) +lean_object* x_161; lean_object* x_162; uint8_t x_163; +x_161 = lean_array_get_size(x_40); +x_162 = lean_unsigned_to_nat(3u); +x_163 = lean_nat_dec_eq(x_161, x_162); +lean_dec(x_161); +if (x_163 == 0) { -lean_object* x_162; -lean_dec(x_38); -x_162 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_162; +lean_object* x_164; +lean_dec(x_40); +x_164 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_164; } else { -lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; -x_163 = lean_unsigned_to_nat(0u); -x_164 = lean_array_fget(x_38, x_163); -x_165 = lean_unsigned_to_nat(2u); -x_166 = lean_array_fget(x_38, x_165); -lean_dec(x_38); -if (lean_obj_tag(x_166) == 5) -{ -lean_object* x_167; -x_167 = lean_ctor_get(x_166, 0); -lean_inc(x_167); -if (lean_obj_tag(x_167) == 5) -{ -lean_object* x_168; -x_168 = lean_ctor_get(x_167, 0); -lean_inc(x_168); +lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; +x_165 = lean_unsigned_to_nat(0u); +x_166 = lean_array_fget(x_40, x_165); +x_167 = lean_unsigned_to_nat(2u); +x_168 = lean_array_fget(x_40, x_167); +lean_dec(x_40); if (lean_obj_tag(x_168) == 5) { lean_object* x_169; x_169 = lean_ctor_get(x_168, 0); lean_inc(x_169); -lean_dec(x_168); if (lean_obj_tag(x_169) == 5) { lean_object* x_170; x_170 = lean_ctor_get(x_169, 0); lean_inc(x_170); -lean_dec(x_169); -if (lean_obj_tag(x_170) == 4) +if (lean_obj_tag(x_170) == 5) { lean_object* x_171; x_171 = lean_ctor_get(x_170, 0); lean_inc(x_171); -if (lean_obj_tag(x_171) == 1) +lean_dec(x_170); +if (lean_obj_tag(x_171) == 5) { lean_object* x_172; x_172 = lean_ctor_get(x_171, 0); lean_inc(x_172); -if (lean_obj_tag(x_172) == 1) +lean_dec(x_171); +if (lean_obj_tag(x_172) == 4) { lean_object* x_173; x_173 = lean_ctor_get(x_172, 0); lean_inc(x_173); -if (lean_obj_tag(x_173) == 0) +if (lean_obj_tag(x_173) == 1) { -lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; uint8_t x_179; -x_174 = lean_ctor_get(x_166, 1); +lean_object* x_174; +x_174 = lean_ctor_get(x_173, 0); lean_inc(x_174); -lean_dec(x_166); -x_175 = lean_ctor_get(x_167, 1); +if (lean_obj_tag(x_174) == 1) +{ +lean_object* x_175; +x_175 = lean_ctor_get(x_174, 0); lean_inc(x_175); -lean_dec(x_167); -x_176 = lean_ctor_get(x_170, 1); +if (lean_obj_tag(x_175) == 0) +{ +lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; uint8_t x_181; +x_176 = lean_ctor_get(x_168, 1); lean_inc(x_176); -lean_dec(x_170); -x_177 = lean_ctor_get(x_171, 1); +lean_dec(x_168); +x_177 = lean_ctor_get(x_169, 1); lean_inc(x_177); -lean_dec(x_171); +lean_dec(x_169); x_178 = lean_ctor_get(x_172, 1); lean_inc(x_178); lean_dec(x_172); -x_179 = lean_string_dec_eq(x_178, x_57); -lean_dec(x_178); -if (x_179 == 0) -{ -lean_object* x_180; -lean_dec(x_177); -lean_dec(x_176); -lean_dec(x_175); +x_179 = lean_ctor_get(x_173, 1); +lean_inc(x_179); +lean_dec(x_173); +x_180 = lean_ctor_get(x_174, 1); +lean_inc(x_180); lean_dec(x_174); -lean_dec(x_164); -x_180 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_180; -} -else +x_181 = lean_string_dec_eq(x_180, x_59); +lean_dec(x_180); +if (x_181 == 0) { -lean_object* x_181; uint8_t x_182; -x_181 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__5; -x_182 = lean_string_dec_eq(x_177, x_181); +lean_object* x_182; +lean_dec(x_179); +lean_dec(x_178); lean_dec(x_177); -if (x_182 == 0) -{ -lean_object* x_183; lean_dec(x_176); -lean_dec(x_175); -lean_dec(x_174); -lean_dec(x_164); -x_183 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_183; +lean_dec(x_166); +x_182 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_182; } else { -if (lean_obj_tag(x_176) == 0) +lean_object* x_183; uint8_t x_184; +x_183 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__5; +x_184 = lean_string_dec_eq(x_179, x_183); +lean_dec(x_179); +if (x_184 == 0) { -lean_object* x_184; -lean_dec(x_175); -lean_dec(x_174); -lean_dec(x_164); -x_184 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_184; +lean_object* x_185; +lean_dec(x_178); +lean_dec(x_177); +lean_dec(x_176); +lean_dec(x_166); +x_185 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_185; } else { -lean_object* x_185; -x_185 = lean_ctor_get(x_176, 1); -lean_inc(x_185); -if (lean_obj_tag(x_185) == 0) +if (lean_obj_tag(x_178) == 0) { lean_object* x_186; +lean_dec(x_177); lean_dec(x_176); -lean_dec(x_175); -lean_dec(x_174); -lean_dec(x_164); -x_186 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_166); +x_186 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); return x_186; } else { lean_object* x_187; -x_187 = lean_ctor_get(x_185, 0); +x_187 = lean_ctor_get(x_178, 1); lean_inc(x_187); if (lean_obj_tag(x_187) == 0) { -uint8_t x_188; -x_188 = !lean_is_exclusive(x_185); -if (x_188 == 0) +lean_object* x_188; +lean_dec(x_178); +lean_dec(x_177); +lean_dec(x_176); +lean_dec(x_166); +x_188 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_188; +} +else { -lean_object* x_189; lean_object* x_190; -x_189 = lean_ctor_get(x_185, 1); -x_190 = lean_ctor_get(x_185, 0); -lean_dec(x_190); +lean_object* x_189; +x_189 = lean_ctor_get(x_187, 0); +lean_inc(x_189); if (lean_obj_tag(x_189) == 0) { -lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; -x_191 = lean_ctor_get(x_176, 0); -lean_inc(x_191); -lean_dec(x_176); -x_192 = lean_box(0); -lean_ctor_set(x_185, 1, x_192); -lean_ctor_set(x_185, 0, x_191); -x_193 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__25; -x_194 = l_Lean_Expr_const___override(x_193, x_185); -x_195 = l_Lean_mkApp3(x_194, x_164, x_175, x_174); -x_196 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_195, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_196; +uint8_t x_190; +x_190 = !lean_is_exclusive(x_187); +if (x_190 == 0) +{ +lean_object* x_191; lean_object* x_192; +x_191 = lean_ctor_get(x_187, 1); +x_192 = lean_ctor_get(x_187, 0); +lean_dec(x_192); +if (lean_obj_tag(x_191) == 0) +{ +lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; +x_193 = lean_ctor_get(x_178, 0); +lean_inc(x_193); +lean_dec(x_178); +x_194 = lean_box(0); +lean_ctor_set(x_187, 1, x_194); +lean_ctor_set(x_187, 0, x_193); +x_195 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__25; +x_196 = l_Lean_Expr_const___override(x_195, x_187); +x_197 = l_Lean_mkApp3(x_196, x_166, x_177, x_176); +x_198 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_197, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_198; } else { -lean_object* x_197; -lean_free_object(x_185); -lean_dec(x_189); +lean_object* x_199; +lean_free_object(x_187); +lean_dec(x_191); +lean_dec(x_178); +lean_dec(x_177); lean_dec(x_176); -lean_dec(x_175); -lean_dec(x_174); -lean_dec(x_164); -x_197 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_197; +lean_dec(x_166); +x_199 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_199; } } else { -lean_object* x_198; -x_198 = lean_ctor_get(x_185, 1); -lean_inc(x_198); -lean_dec(x_185); -if (lean_obj_tag(x_198) == 0) +lean_object* x_200; +x_200 = lean_ctor_get(x_187, 1); +lean_inc(x_200); +lean_dec(x_187); +if (lean_obj_tag(x_200) == 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; lean_object* x_205; -x_199 = lean_ctor_get(x_176, 0); -lean_inc(x_199); -lean_dec(x_176); -x_200 = lean_box(0); -x_201 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_201, 0, x_199); -lean_ctor_set(x_201, 1, x_200); -x_202 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__25; -x_203 = l_Lean_Expr_const___override(x_202, x_201); -x_204 = l_Lean_mkApp3(x_203, x_164, x_175, x_174); -x_205 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_204, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_205; +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; +x_201 = lean_ctor_get(x_178, 0); +lean_inc(x_201); +lean_dec(x_178); +x_202 = lean_box(0); +x_203 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_203, 0, x_201); +lean_ctor_set(x_203, 1, x_202); +x_204 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__25; +x_205 = l_Lean_Expr_const___override(x_204, x_203); +x_206 = l_Lean_mkApp3(x_205, x_166, x_177, x_176); +x_207 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_206, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_207; } else { -lean_object* x_206; -lean_dec(x_198); +lean_object* x_208; +lean_dec(x_200); +lean_dec(x_178); +lean_dec(x_177); lean_dec(x_176); -lean_dec(x_175); -lean_dec(x_174); -lean_dec(x_164); -x_206 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_206; +lean_dec(x_166); +x_208 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_208; } } } else { -lean_object* x_207; +lean_object* x_209; +lean_dec(x_189); lean_dec(x_187); -lean_dec(x_185); +lean_dec(x_178); +lean_dec(x_177); lean_dec(x_176); -lean_dec(x_175); -lean_dec(x_174); -lean_dec(x_164); -x_207 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_207; +lean_dec(x_166); +x_209 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_209; } } } @@ -13367,451 +13479,451 @@ return x_207; } else { -lean_object* x_208; +lean_object* x_210; +lean_dec(x_175); +lean_dec(x_174); lean_dec(x_173); lean_dec(x_172); -lean_dec(x_171); -lean_dec(x_170); -lean_dec(x_167); +lean_dec(x_169); +lean_dec(x_168); lean_dec(x_166); -lean_dec(x_164); -x_208 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_208; +x_210 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_210; } } else { -lean_object* x_209; +lean_object* x_211; +lean_dec(x_174); +lean_dec(x_173); lean_dec(x_172); -lean_dec(x_171); -lean_dec(x_170); -lean_dec(x_167); +lean_dec(x_169); +lean_dec(x_168); lean_dec(x_166); -lean_dec(x_164); -x_209 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_209; +x_211 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_211; } } else { -lean_object* x_210; -lean_dec(x_171); -lean_dec(x_170); -lean_dec(x_167); +lean_object* x_212; +lean_dec(x_173); +lean_dec(x_172); +lean_dec(x_169); +lean_dec(x_168); lean_dec(x_166); -lean_dec(x_164); -x_210 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_210; +x_212 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_212; } } else { -lean_object* x_211; -lean_dec(x_170); -lean_dec(x_167); +lean_object* x_213; +lean_dec(x_172); +lean_dec(x_169); +lean_dec(x_168); lean_dec(x_166); -lean_dec(x_164); -x_211 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_211; +x_213 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_213; } } else { -lean_object* x_212; +lean_object* x_214; +lean_dec(x_171); lean_dec(x_169); -lean_dec(x_167); +lean_dec(x_168); lean_dec(x_166); -lean_dec(x_164); -x_212 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_212; +x_214 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_214; } } else { -lean_object* x_213; +lean_object* x_215; +lean_dec(x_170); +lean_dec(x_169); lean_dec(x_168); -lean_dec(x_167); lean_dec(x_166); -lean_dec(x_164); -x_213 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_213; +x_215 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_215; } } else { -lean_object* x_214; -lean_dec(x_167); +lean_object* x_216; +lean_dec(x_169); +lean_dec(x_168); lean_dec(x_166); -lean_dec(x_164); -x_214 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_214; +x_216 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_216; } } else { -lean_object* x_215; +lean_object* x_217; +lean_dec(x_168); lean_dec(x_166); -lean_dec(x_164); -x_215 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_215; +x_217 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_217; } } } } else { -lean_object* x_216; lean_object* x_217; uint8_t x_218; -lean_dec(x_39); -x_216 = lean_array_get_size(x_38); -x_217 = lean_unsigned_to_nat(3u); -x_218 = lean_nat_dec_eq(x_216, x_217); -lean_dec(x_216); -if (x_218 == 0) +lean_object* x_218; lean_object* x_219; uint8_t x_220; +lean_dec(x_41); +x_218 = lean_array_get_size(x_40); +x_219 = lean_unsigned_to_nat(3u); +x_220 = lean_nat_dec_eq(x_218, x_219); +lean_dec(x_218); +if (x_220 == 0) { -lean_object* x_219; -lean_dec(x_38); -x_219 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_219; +lean_object* x_221; +lean_dec(x_40); +x_221 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_221; } else { -lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; -x_220 = lean_unsigned_to_nat(1u); -x_221 = lean_array_fget(x_38, x_220); -x_222 = lean_unsigned_to_nat(2u); -x_223 = lean_array_fget(x_38, x_222); -lean_dec(x_38); -if (lean_obj_tag(x_223) == 5) -{ -lean_object* x_224; -x_224 = lean_ctor_get(x_223, 0); -lean_inc(x_224); -if (lean_obj_tag(x_224) == 5) -{ -lean_object* x_225; -x_225 = lean_ctor_get(x_224, 0); -lean_inc(x_225); +lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; +x_222 = lean_unsigned_to_nat(1u); +x_223 = lean_array_fget(x_40, x_222); +x_224 = lean_unsigned_to_nat(2u); +x_225 = lean_array_fget(x_40, x_224); +lean_dec(x_40); if (lean_obj_tag(x_225) == 5) { lean_object* x_226; x_226 = lean_ctor_get(x_225, 0); lean_inc(x_226); -lean_dec(x_225); if (lean_obj_tag(x_226) == 5) { lean_object* x_227; x_227 = lean_ctor_get(x_226, 0); lean_inc(x_227); -lean_dec(x_226); -if (lean_obj_tag(x_227) == 4) +if (lean_obj_tag(x_227) == 5) { lean_object* x_228; x_228 = lean_ctor_get(x_227, 0); lean_inc(x_228); -if (lean_obj_tag(x_228) == 1) +lean_dec(x_227); +if (lean_obj_tag(x_228) == 5) { lean_object* x_229; x_229 = lean_ctor_get(x_228, 0); lean_inc(x_229); -if (lean_obj_tag(x_229) == 1) +lean_dec(x_228); +if (lean_obj_tag(x_229) == 4) { lean_object* x_230; x_230 = lean_ctor_get(x_229, 0); lean_inc(x_230); -if (lean_obj_tag(x_230) == 0) +if (lean_obj_tag(x_230) == 1) { -lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; uint8_t x_236; -x_231 = lean_ctor_get(x_223, 1); +lean_object* x_231; +x_231 = lean_ctor_get(x_230, 0); lean_inc(x_231); -lean_dec(x_223); -x_232 = lean_ctor_get(x_224, 1); +if (lean_obj_tag(x_231) == 1) +{ +lean_object* x_232; +x_232 = lean_ctor_get(x_231, 0); lean_inc(x_232); -lean_dec(x_224); -x_233 = lean_ctor_get(x_227, 1); +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; +x_233 = lean_ctor_get(x_225, 1); lean_inc(x_233); -lean_dec(x_227); -x_234 = lean_ctor_get(x_228, 1); +lean_dec(x_225); +x_234 = lean_ctor_get(x_226, 1); lean_inc(x_234); -lean_dec(x_228); +lean_dec(x_226); x_235 = lean_ctor_get(x_229, 1); lean_inc(x_235); lean_dec(x_229); -x_236 = lean_string_dec_eq(x_235, x_57); -lean_dec(x_235); -if (x_236 == 0) -{ -lean_object* x_237; -lean_dec(x_234); -lean_dec(x_233); -lean_dec(x_232); +x_236 = lean_ctor_get(x_230, 1); +lean_inc(x_236); +lean_dec(x_230); +x_237 = lean_ctor_get(x_231, 1); +lean_inc(x_237); lean_dec(x_231); -lean_dec(x_221); -x_237 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_237; -} -else +x_238 = lean_string_dec_eq(x_237, x_59); +lean_dec(x_237); +if (x_238 == 0) { -lean_object* x_238; uint8_t x_239; -x_238 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__5; -x_239 = lean_string_dec_eq(x_234, x_238); +lean_object* x_239; +lean_dec(x_236); +lean_dec(x_235); lean_dec(x_234); -if (x_239 == 0) -{ -lean_object* x_240; lean_dec(x_233); -lean_dec(x_232); -lean_dec(x_231); -lean_dec(x_221); -x_240 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_240; +lean_dec(x_223); +x_239 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_239; } else { -if (lean_obj_tag(x_233) == 0) +lean_object* x_240; uint8_t x_241; +x_240 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__5; +x_241 = lean_string_dec_eq(x_236, x_240); +lean_dec(x_236); +if (x_241 == 0) { -lean_object* x_241; -lean_dec(x_232); -lean_dec(x_231); -lean_dec(x_221); -x_241 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_241; +lean_object* x_242; +lean_dec(x_235); +lean_dec(x_234); +lean_dec(x_233); +lean_dec(x_223); +x_242 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_242; } else { -lean_object* x_242; -x_242 = lean_ctor_get(x_233, 0); -lean_inc(x_242); -if (lean_obj_tag(x_242) == 0) +if (lean_obj_tag(x_235) == 0) { lean_object* x_243; -x_243 = lean_ctor_get(x_233, 1); -lean_inc(x_243); +lean_dec(x_234); lean_dec(x_233); -if (lean_obj_tag(x_243) == 0) -{ -lean_object* x_244; -lean_dec(x_232); -lean_dec(x_231); -lean_dec(x_221); -x_244 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_244; +lean_dec(x_223); +x_243 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_243; } else { +lean_object* x_244; +x_244 = lean_ctor_get(x_235, 0); +lean_inc(x_244); +if (lean_obj_tag(x_244) == 0) +{ lean_object* x_245; -x_245 = lean_ctor_get(x_243, 1); +x_245 = lean_ctor_get(x_235, 1); lean_inc(x_245); +lean_dec(x_235); if (lean_obj_tag(x_245) == 0) { -uint8_t x_246; -x_246 = !lean_is_exclusive(x_243); -if (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; -x_247 = lean_ctor_get(x_243, 1); -lean_dec(x_247); -x_248 = lean_box(0); -lean_ctor_set(x_243, 1, x_248); -x_249 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__27; -x_250 = l_Lean_Expr_const___override(x_249, x_243); -x_251 = l_Lean_mkApp3(x_250, x_221, x_232, x_231); -x_252 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_251, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_252; +lean_object* x_246; +lean_dec(x_234); +lean_dec(x_233); +lean_dec(x_223); +x_246 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_246; } else { -lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; -x_253 = lean_ctor_get(x_243, 0); -lean_inc(x_253); -lean_dec(x_243); -x_254 = lean_box(0); -x_255 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_255, 0, x_253); -lean_ctor_set(x_255, 1, x_254); -x_256 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__27; -x_257 = l_Lean_Expr_const___override(x_256, x_255); -x_258 = l_Lean_mkApp3(x_257, x_221, x_232, x_231); -x_259 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_258, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_259; -} -} -else +lean_object* x_247; +x_247 = lean_ctor_get(x_245, 1); +lean_inc(x_247); +if (lean_obj_tag(x_247) == 0) { -lean_object* x_260; -lean_dec(x_245); -lean_dec(x_243); -lean_dec(x_232); -lean_dec(x_231); -lean_dec(x_221); -x_260 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_260; -} -} +uint8_t x_248; +x_248 = !lean_is_exclusive(x_245); +if (x_248 == 0) +{ +lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; +x_249 = lean_ctor_get(x_245, 1); +lean_dec(x_249); +x_250 = lean_box(0); +lean_ctor_set(x_245, 1, x_250); +x_251 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__27; +x_252 = l_Lean_Expr_const___override(x_251, x_245); +x_253 = l_Lean_mkApp3(x_252, x_223, x_234, x_233); +x_254 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_253, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_254; } else { -lean_object* x_261; -lean_dec(x_242); -lean_dec(x_233); -lean_dec(x_232); -lean_dec(x_231); -lean_dec(x_221); -x_261 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +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_255 = lean_ctor_get(x_245, 0); +lean_inc(x_255); +lean_dec(x_245); +x_256 = lean_box(0); +x_257 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_257, 0, x_255); +lean_ctor_set(x_257, 1, x_256); +x_258 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__27; +x_259 = l_Lean_Expr_const___override(x_258, x_257); +x_260 = l_Lean_mkApp3(x_259, x_223, x_234, x_233); +x_261 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_260, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); return x_261; } } -} -} -} else { lean_object* x_262; -lean_dec(x_230); -lean_dec(x_229); -lean_dec(x_228); -lean_dec(x_227); -lean_dec(x_224); +lean_dec(x_247); +lean_dec(x_245); +lean_dec(x_234); +lean_dec(x_233); lean_dec(x_223); -lean_dec(x_221); -x_262 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_262 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); return x_262; } } +} else { lean_object* x_263; -lean_dec(x_229); -lean_dec(x_228); -lean_dec(x_227); -lean_dec(x_224); +lean_dec(x_244); +lean_dec(x_235); +lean_dec(x_234); +lean_dec(x_233); lean_dec(x_223); -lean_dec(x_221); -x_263 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_263 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); return x_263; } } +} +} +} else { lean_object* x_264; -lean_dec(x_228); -lean_dec(x_227); -lean_dec(x_224); +lean_dec(x_232); +lean_dec(x_231); +lean_dec(x_230); +lean_dec(x_229); +lean_dec(x_226); +lean_dec(x_225); lean_dec(x_223); -lean_dec(x_221); -x_264 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_264 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); return x_264; } } else { lean_object* x_265; -lean_dec(x_227); -lean_dec(x_224); +lean_dec(x_231); +lean_dec(x_230); +lean_dec(x_229); +lean_dec(x_226); +lean_dec(x_225); lean_dec(x_223); -lean_dec(x_221); -x_265 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_265 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); return x_265; } } else { lean_object* x_266; +lean_dec(x_230); +lean_dec(x_229); lean_dec(x_226); -lean_dec(x_224); +lean_dec(x_225); lean_dec(x_223); -lean_dec(x_221); -x_266 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_266 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); return x_266; } } else { lean_object* x_267; +lean_dec(x_229); +lean_dec(x_226); lean_dec(x_225); -lean_dec(x_224); lean_dec(x_223); -lean_dec(x_221); -x_267 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_267 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); return x_267; } } else { lean_object* x_268; -lean_dec(x_224); +lean_dec(x_228); +lean_dec(x_226); +lean_dec(x_225); lean_dec(x_223); -lean_dec(x_221); -x_268 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_268 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); return x_268; } } else { lean_object* x_269; +lean_dec(x_227); +lean_dec(x_226); +lean_dec(x_225); lean_dec(x_223); -lean_dec(x_221); -x_269 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_269 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); return x_269; } } +else +{ +lean_object* x_270; +lean_dec(x_226); +lean_dec(x_225); +lean_dec(x_223); +x_270 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_270; +} +} +else +{ +lean_object* x_271; +lean_dec(x_225); +lean_dec(x_223); +x_271 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_271; +} +} } } } else { -lean_object* x_270; uint8_t x_271; -lean_dec(x_40); +lean_object* x_272; uint8_t x_273; +lean_dec(x_42); lean_dec(x_2); -x_270 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__16; -x_271 = lean_string_dec_eq(x_39, x_270); -lean_dec(x_39); -if (x_271 == 0) +x_272 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__16; +x_273 = lean_string_dec_eq(x_41, x_272); +lean_dec(x_41); +if (x_273 == 0) { -lean_object* x_272; -lean_dec(x_38); -x_272 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_272; +lean_object* x_274; +lean_dec(x_40); +x_274 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_274; } else { -lean_object* x_273; lean_object* x_274; uint8_t x_275; -x_273 = lean_array_get_size(x_38); -x_274 = lean_unsigned_to_nat(6u); -x_275 = lean_nat_dec_eq(x_273, x_274); -lean_dec(x_273); -if (x_275 == 0) +lean_object* x_275; lean_object* x_276; uint8_t x_277; +x_275 = lean_array_get_size(x_40); +x_276 = lean_unsigned_to_nat(6u); +x_277 = lean_nat_dec_eq(x_275, x_276); +lean_dec(x_275); +if (x_277 == 0) { -lean_object* x_276; -lean_dec(x_38); -x_276 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_276; +lean_object* x_278; +lean_dec(x_40); +x_278 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_278; } else { -lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; -x_277 = lean_unsigned_to_nat(4u); -x_278 = lean_array_fget(x_38, x_277); -x_279 = lean_unsigned_to_nat(5u); -x_280 = lean_array_fget(x_38, x_279); -lean_dec(x_38); -lean_inc(x_278); -x_281 = l_Lean_Elab_Tactic_Omega_groundNat_x3f(x_278); -if (lean_obj_tag(x_281) == 0) +lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; +x_279 = lean_unsigned_to_nat(4u); +x_280 = lean_array_fget(x_40, x_279); +x_281 = lean_unsigned_to_nat(5u); +x_282 = lean_array_fget(x_40, x_281); +lean_dec(x_40); +lean_inc(x_280); +x_283 = l_Lean_Elab_Tactic_Omega_groundNat_x3f(x_280); +if (lean_obj_tag(x_283) == 0) { -lean_object* x_282; +lean_object* x_284; +lean_dec(x_282); lean_dec(x_280); -lean_dec(x_278); -x_282 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_282; +x_284 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_284; } else { -lean_object* x_283; lean_object* x_284; lean_object* x_285; -lean_dec(x_281); -x_283 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__30; -x_284 = l_Lean_mkAppB(x_283, x_278, x_280); -x_285 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_284, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_285; +lean_object* x_285; lean_object* x_286; lean_object* x_287; +lean_dec(x_283); +x_285 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__30; +x_286 = l_Lean_mkAppB(x_285, x_280, x_282); +x_287 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_286, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_287; } } } @@ -13819,60 +13931,48 @@ return x_285; } else { -lean_object* x_286; uint8_t x_287; -lean_dec(x_40); +lean_object* x_288; uint8_t x_289; +lean_dec(x_42); lean_dec(x_2); -x_286 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__47; -x_287 = lean_string_dec_eq(x_39, x_286); -lean_dec(x_39); -if (x_287 == 0) +x_288 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__47; +x_289 = lean_string_dec_eq(x_41, x_288); +lean_dec(x_41); +if (x_289 == 0) { -lean_object* x_288; -lean_dec(x_38); -x_288 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_288; +lean_object* x_290; +lean_dec(x_40); +x_290 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_290; } else { -lean_object* x_289; lean_object* x_290; uint8_t x_291; -x_289 = lean_array_get_size(x_38); -x_290 = lean_unsigned_to_nat(6u); -x_291 = lean_nat_dec_eq(x_289, x_290); -lean_dec(x_289); -if (x_291 == 0) +lean_object* x_291; lean_object* x_292; uint8_t x_293; +x_291 = lean_array_get_size(x_40); +x_292 = lean_unsigned_to_nat(6u); +x_293 = lean_nat_dec_eq(x_291, x_292); +lean_dec(x_291); +if (x_293 == 0) { -lean_object* x_292; -lean_dec(x_38); -x_292 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_292; +lean_object* x_294; +lean_dec(x_40); +x_294 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_294; } else { -lean_object* x_293; lean_object* x_294; -x_293 = lean_unsigned_to_nat(4u); -x_294 = lean_array_fget(x_38, x_293); -if (lean_obj_tag(x_294) == 5) -{ -lean_object* x_295; -x_295 = lean_ctor_get(x_294, 0); -lean_inc(x_295); -if (lean_obj_tag(x_295) == 5) -{ -lean_object* x_296; -x_296 = lean_ctor_get(x_295, 0); -lean_inc(x_296); +lean_object* x_295; lean_object* x_296; +x_295 = lean_unsigned_to_nat(4u); +x_296 = lean_array_fget(x_40, x_295); if (lean_obj_tag(x_296) == 5) { lean_object* x_297; x_297 = lean_ctor_get(x_296, 0); lean_inc(x_297); -lean_dec(x_296); if (lean_obj_tag(x_297) == 5) { lean_object* x_298; x_298 = lean_ctor_get(x_297, 0); lean_inc(x_298); -lean_dec(x_297); if (lean_obj_tag(x_298) == 5) { lean_object* x_299; @@ -13885,486 +13985,498 @@ lean_object* x_300; x_300 = lean_ctor_get(x_299, 0); lean_inc(x_300); lean_dec(x_299); -if (lean_obj_tag(x_300) == 4) +if (lean_obj_tag(x_300) == 5) { lean_object* x_301; x_301 = lean_ctor_get(x_300, 0); lean_inc(x_301); lean_dec(x_300); -if (lean_obj_tag(x_301) == 1) +if (lean_obj_tag(x_301) == 5) { lean_object* x_302; x_302 = lean_ctor_get(x_301, 0); lean_inc(x_302); -if (lean_obj_tag(x_302) == 1) +lean_dec(x_301); +if (lean_obj_tag(x_302) == 4) { lean_object* x_303; x_303 = lean_ctor_get(x_302, 0); lean_inc(x_303); -if (lean_obj_tag(x_303) == 0) +lean_dec(x_302); +if (lean_obj_tag(x_303) == 1) { -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_294, 1); +lean_object* x_304; +x_304 = lean_ctor_get(x_303, 0); lean_inc(x_304); -lean_dec(x_294); -x_305 = lean_ctor_get(x_295, 1); +if (lean_obj_tag(x_304) == 1) +{ +lean_object* x_305; +x_305 = lean_ctor_get(x_304, 0); lean_inc(x_305); -lean_dec(x_295); -x_306 = lean_ctor_get(x_301, 1); +if (lean_obj_tag(x_305) == 0) +{ +lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; uint8_t x_310; +x_306 = lean_ctor_get(x_296, 1); lean_inc(x_306); -lean_dec(x_301); -x_307 = lean_ctor_get(x_302, 1); +lean_dec(x_296); +x_307 = lean_ctor_get(x_297, 1); lean_inc(x_307); -lean_dec(x_302); -x_308 = lean_string_dec_eq(x_307, x_53); -lean_dec(x_307); -if (x_308 == 0) +lean_dec(x_297); +x_308 = lean_ctor_get(x_303, 1); +lean_inc(x_308); +lean_dec(x_303); +x_309 = lean_ctor_get(x_304, 1); +lean_inc(x_309); +lean_dec(x_304); +x_310 = lean_string_dec_eq(x_309, x_55); +lean_dec(x_309); +if (x_310 == 0) { -lean_object* x_309; +lean_object* x_311; +lean_dec(x_308); +lean_dec(x_307); lean_dec(x_306); -lean_dec(x_305); -lean_dec(x_304); -lean_dec(x_38); -x_309 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_309; +lean_dec(x_40); +x_311 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_311; } else { -uint8_t x_310; -x_310 = lean_string_dec_eq(x_306, x_286); -lean_dec(x_306); -if (x_310 == 0) -{ -lean_object* x_311; -lean_dec(x_305); -lean_dec(x_304); -lean_dec(x_38); -x_311 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_311; -} -else -{ -lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; -x_312 = lean_unsigned_to_nat(5u); -x_313 = lean_array_fget(x_38, x_312); -lean_dec(x_38); -x_314 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__33; -x_315 = l_Lean_mkApp3(x_314, x_305, x_304, x_313); -x_316 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_315, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_316; -} -} -} -else +uint8_t x_312; +x_312 = lean_string_dec_eq(x_308, x_288); +lean_dec(x_308); +if (x_312 == 0) { -lean_object* x_317; -lean_dec(x_303); -lean_dec(x_302); -lean_dec(x_301); -lean_dec(x_295); -lean_dec(x_294); -lean_dec(x_38); -x_317 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_317; -} +lean_object* x_313; +lean_dec(x_307); +lean_dec(x_306); +lean_dec(x_40); +x_313 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_313; } else { -lean_object* x_318; -lean_dec(x_302); -lean_dec(x_301); -lean_dec(x_295); -lean_dec(x_294); -lean_dec(x_38); -x_318 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; +x_314 = lean_unsigned_to_nat(5u); +x_315 = lean_array_fget(x_40, x_314); +lean_dec(x_40); +x_316 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__33; +x_317 = l_Lean_mkApp3(x_316, x_307, x_306, x_315); +x_318 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_317, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); return x_318; } } +} else { lean_object* x_319; -lean_dec(x_301); -lean_dec(x_295); -lean_dec(x_294); -lean_dec(x_38); -x_319 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_305); +lean_dec(x_304); +lean_dec(x_303); +lean_dec(x_297); +lean_dec(x_296); +lean_dec(x_40); +x_319 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); return x_319; } } else { lean_object* x_320; -lean_dec(x_300); -lean_dec(x_295); -lean_dec(x_294); -lean_dec(x_38); -x_320 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_304); +lean_dec(x_303); +lean_dec(x_297); +lean_dec(x_296); +lean_dec(x_40); +x_320 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); return x_320; } } else { lean_object* x_321; -lean_dec(x_299); -lean_dec(x_295); -lean_dec(x_294); -lean_dec(x_38); -x_321 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_303); +lean_dec(x_297); +lean_dec(x_296); +lean_dec(x_40); +x_321 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); return x_321; } } else { lean_object* x_322; -lean_dec(x_298); -lean_dec(x_295); -lean_dec(x_294); -lean_dec(x_38); -x_322 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_302); +lean_dec(x_297); +lean_dec(x_296); +lean_dec(x_40); +x_322 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); return x_322; } } else { lean_object* x_323; +lean_dec(x_301); lean_dec(x_297); -lean_dec(x_295); -lean_dec(x_294); -lean_dec(x_38); -x_323 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_296); +lean_dec(x_40); +x_323 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); return x_323; } } else { lean_object* x_324; +lean_dec(x_300); +lean_dec(x_297); lean_dec(x_296); -lean_dec(x_295); -lean_dec(x_294); -lean_dec(x_38); -x_324 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_40); +x_324 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); return x_324; } } else { lean_object* x_325; -lean_dec(x_295); -lean_dec(x_294); -lean_dec(x_38); -x_325 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_299); +lean_dec(x_297); +lean_dec(x_296); +lean_dec(x_40); +x_325 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); return x_325; } } else { lean_object* x_326; -lean_dec(x_294); -lean_dec(x_38); -x_326 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_298); +lean_dec(x_297); +lean_dec(x_296); +lean_dec(x_40); +x_326 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); return x_326; } } -} +else +{ +lean_object* x_327; +lean_dec(x_297); +lean_dec(x_296); +lean_dec(x_40); +x_327 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_327; } } else { -lean_object* x_327; uint8_t x_328; +lean_object* x_328; +lean_dec(x_296); lean_dec(x_40); +x_328 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_328; +} +} +} +} +} +else +{ +lean_object* x_329; uint8_t x_330; +lean_dec(x_42); lean_dec(x_2); -x_327 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__39; -x_328 = lean_string_dec_eq(x_39, x_327); -lean_dec(x_39); -if (x_328 == 0) +x_329 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__39; +x_330 = lean_string_dec_eq(x_41, x_329); +lean_dec(x_41); +if (x_330 == 0) { -lean_object* x_329; -lean_dec(x_38); -x_329 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_329; +lean_object* x_331; +lean_dec(x_40); +x_331 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_331; } else { -lean_object* x_330; lean_object* x_331; uint8_t x_332; -x_330 = lean_array_get_size(x_38); -x_331 = lean_unsigned_to_nat(6u); -x_332 = lean_nat_dec_eq(x_330, x_331); -lean_dec(x_330); -if (x_332 == 0) +lean_object* x_332; lean_object* x_333; uint8_t x_334; +x_332 = lean_array_get_size(x_40); +x_333 = lean_unsigned_to_nat(6u); +x_334 = lean_nat_dec_eq(x_332, x_333); +lean_dec(x_332); +if (x_334 == 0) { -lean_object* x_333; -lean_dec(x_38); -x_333 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_333; +lean_object* x_335; +lean_dec(x_40); +x_335 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_335; } else { -lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; -x_334 = lean_unsigned_to_nat(4u); -x_335 = lean_array_fget(x_38, x_334); -x_336 = lean_unsigned_to_nat(5u); -x_337 = lean_array_fget(x_38, x_336); -lean_dec(x_38); -x_338 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__36; -x_339 = l_Lean_mkAppB(x_338, x_335, x_337); -x_340 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_339, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_340; +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; +x_336 = lean_unsigned_to_nat(4u); +x_337 = lean_array_fget(x_40, x_336); +x_338 = lean_unsigned_to_nat(5u); +x_339 = lean_array_fget(x_40, x_338); +lean_dec(x_40); +x_340 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__36; +x_341 = l_Lean_mkAppB(x_340, x_337, x_339); +x_342 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_341, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_342; } } } } else { -lean_object* x_341; uint8_t x_342; -lean_dec(x_40); +lean_object* x_343; uint8_t x_344; +lean_dec(x_42); lean_dec(x_2); -x_341 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__37; -x_342 = lean_string_dec_eq(x_39, x_341); -lean_dec(x_39); -if (x_342 == 0) +x_343 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__37; +x_344 = lean_string_dec_eq(x_41, x_343); +lean_dec(x_41); +if (x_344 == 0) { -lean_object* x_343; -lean_dec(x_38); -x_343 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_343; +lean_object* x_345; +lean_dec(x_40); +x_345 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_345; } else { -lean_object* x_344; lean_object* x_345; uint8_t x_346; -x_344 = lean_array_get_size(x_38); -x_345 = lean_unsigned_to_nat(3u); -x_346 = lean_nat_dec_eq(x_344, x_345); -lean_dec(x_344); -if (x_346 == 0) +lean_object* x_346; lean_object* x_347; uint8_t x_348; +x_346 = lean_array_get_size(x_40); +x_347 = lean_unsigned_to_nat(3u); +x_348 = lean_nat_dec_eq(x_346, x_347); +lean_dec(x_346); +if (x_348 == 0) { -lean_object* x_347; -lean_dec(x_38); -x_347 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_347; +lean_object* x_349; +lean_dec(x_40); +x_349 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_349; } else { -lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; -x_348 = lean_unsigned_to_nat(1u); -x_349 = lean_array_fget(x_38, x_348); -lean_dec(x_38); -x_350 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__40; -x_351 = l_Lean_Expr_app___override(x_350, x_349); -x_352 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_351, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_352; +lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; +x_350 = lean_unsigned_to_nat(1u); +x_351 = lean_array_fget(x_40, x_350); +lean_dec(x_40); +x_352 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__40; +x_353 = l_Lean_Expr_app___override(x_352, x_351); +x_354 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_353, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_354; } } } } else { -lean_object* x_353; uint8_t x_354; -lean_dec(x_40); +lean_object* x_355; uint8_t x_356; +lean_dec(x_42); lean_dec(x_2); -x_353 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__28; -x_354 = lean_string_dec_eq(x_39, x_353); -lean_dec(x_39); -if (x_354 == 0) +x_355 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__28; +x_356 = lean_string_dec_eq(x_41, x_355); +lean_dec(x_41); +if (x_356 == 0) { -lean_object* x_355; -lean_dec(x_38); -x_355 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_355; +lean_object* x_357; +lean_dec(x_40); +x_357 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_357; } else { -lean_object* x_356; lean_object* x_357; uint8_t x_358; -x_356 = lean_array_get_size(x_38); -x_357 = lean_unsigned_to_nat(6u); -x_358 = lean_nat_dec_eq(x_356, x_357); -lean_dec(x_356); -if (x_358 == 0) +lean_object* x_358; lean_object* x_359; uint8_t x_360; +x_358 = lean_array_get_size(x_40); +x_359 = lean_unsigned_to_nat(6u); +x_360 = lean_nat_dec_eq(x_358, x_359); +lean_dec(x_358); +if (x_360 == 0) { -lean_object* x_359; -lean_dec(x_38); -x_359 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_359; +lean_object* x_361; +lean_dec(x_40); +x_361 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_361; } else { -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; -x_360 = lean_unsigned_to_nat(4u); -x_361 = lean_array_fget(x_38, x_360); -x_362 = lean_unsigned_to_nat(5u); -x_363 = lean_array_fget(x_38, x_362); -lean_dec(x_38); -x_364 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__43; -x_365 = l_Lean_mkAppB(x_364, x_361, x_363); -x_366 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_365, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_366; +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; +x_362 = lean_unsigned_to_nat(4u); +x_363 = lean_array_fget(x_40, x_362); +x_364 = lean_unsigned_to_nat(5u); +x_365 = lean_array_fget(x_40, x_364); +lean_dec(x_40); +x_366 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__43; +x_367 = l_Lean_mkAppB(x_366, x_363, x_365); +x_368 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_367, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_368; } } } } else { -lean_object* x_367; uint8_t x_368; -lean_dec(x_40); +lean_object* x_369; uint8_t x_370; +lean_dec(x_42); lean_dec(x_2); -x_367 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__45; -x_368 = lean_string_dec_eq(x_39, x_367); -lean_dec(x_39); -if (x_368 == 0) +x_369 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__45; +x_370 = lean_string_dec_eq(x_41, x_369); +lean_dec(x_41); +if (x_370 == 0) { -lean_object* x_369; -lean_dec(x_38); -x_369 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_369; +lean_object* x_371; +lean_dec(x_40); +x_371 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_371; } else { -lean_object* x_370; lean_object* x_371; uint8_t x_372; -x_370 = lean_array_get_size(x_38); -x_371 = lean_unsigned_to_nat(6u); -x_372 = lean_nat_dec_eq(x_370, x_371); -lean_dec(x_370); -if (x_372 == 0) +lean_object* x_372; lean_object* x_373; uint8_t x_374; +x_372 = lean_array_get_size(x_40); +x_373 = lean_unsigned_to_nat(6u); +x_374 = lean_nat_dec_eq(x_372, x_373); +lean_dec(x_372); +if (x_374 == 0) { -lean_object* x_373; -lean_dec(x_38); -x_373 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_373; +lean_object* x_375; +lean_dec(x_40); +x_375 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_375; } else { -lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; -x_374 = lean_unsigned_to_nat(4u); -x_375 = lean_array_fget(x_38, x_374); -x_376 = lean_unsigned_to_nat(5u); -x_377 = lean_array_fget(x_38, x_376); -lean_dec(x_38); -x_378 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__46; -x_379 = l_Lean_mkAppB(x_378, x_375, x_377); -x_380 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_379, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_380; +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; +x_376 = lean_unsigned_to_nat(4u); +x_377 = lean_array_fget(x_40, x_376); +x_378 = lean_unsigned_to_nat(5u); +x_379 = lean_array_fget(x_40, x_378); +lean_dec(x_40); +x_380 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__46; +x_381 = l_Lean_mkAppB(x_380, x_377, x_379); +x_382 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_381, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_382; } } } } else { -lean_object* x_381; uint8_t x_382; -lean_dec(x_40); +lean_object* x_383; uint8_t x_384; +lean_dec(x_42); lean_dec(x_2); -x_381 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__3; -x_382 = lean_string_dec_eq(x_39, x_381); -lean_dec(x_39); -if (x_382 == 0) +x_383 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__3; +x_384 = lean_string_dec_eq(x_41, x_383); +lean_dec(x_41); +if (x_384 == 0) { -lean_object* x_383; -lean_dec(x_38); -x_383 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_383; +lean_object* x_385; +lean_dec(x_40); +x_385 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_385; } else { -lean_object* x_384; lean_object* x_385; uint8_t x_386; -x_384 = lean_array_get_size(x_38); -x_385 = lean_unsigned_to_nat(6u); -x_386 = lean_nat_dec_eq(x_384, x_385); -lean_dec(x_384); -if (x_386 == 0) +lean_object* x_386; lean_object* x_387; uint8_t x_388; +x_386 = lean_array_get_size(x_40); +x_387 = lean_unsigned_to_nat(6u); +x_388 = lean_nat_dec_eq(x_386, x_387); +lean_dec(x_386); +if (x_388 == 0) { -lean_object* x_387; -lean_dec(x_38); -x_387 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_387; +lean_object* x_389; +lean_dec(x_40); +x_389 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_389; } else { -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; -x_388 = lean_unsigned_to_nat(4u); -x_389 = lean_array_fget(x_38, x_388); -x_390 = lean_unsigned_to_nat(5u); -x_391 = lean_array_fget(x_38, x_390); -lean_dec(x_38); -x_392 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__49; -x_393 = l_Lean_mkAppB(x_392, x_389, x_391); -x_394 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_393, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_394; +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; +x_390 = lean_unsigned_to_nat(4u); +x_391 = lean_array_fget(x_40, x_390); +x_392 = lean_unsigned_to_nat(5u); +x_393 = lean_array_fget(x_40, x_392); +lean_dec(x_40); +x_394 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__49; +x_395 = l_Lean_mkAppB(x_394, x_391, x_393); +x_396 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_395, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_396; } } } } else { -lean_object* x_395; uint8_t x_396; -lean_dec(x_40); +lean_object* x_397; uint8_t x_398; +lean_dec(x_42); lean_dec(x_2); -x_395 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__5; -x_396 = lean_string_dec_eq(x_39, x_395); -lean_dec(x_39); -if (x_396 == 0) +x_397 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__5; +x_398 = lean_string_dec_eq(x_41, x_397); +lean_dec(x_41); +if (x_398 == 0) { -lean_object* x_397; -lean_dec(x_38); -x_397 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_397; +lean_object* x_399; +lean_dec(x_40); +x_399 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_399; } else { -lean_object* x_398; lean_object* x_399; uint8_t x_400; -x_398 = lean_array_get_size(x_38); -x_399 = lean_unsigned_to_nat(1u); -x_400 = lean_nat_dec_eq(x_398, x_399); -lean_dec(x_398); -if (x_400 == 0) +lean_object* x_400; lean_object* x_401; uint8_t x_402; +x_400 = lean_array_get_size(x_40); +x_401 = lean_unsigned_to_nat(1u); +x_402 = lean_nat_dec_eq(x_400, x_401); +lean_dec(x_400); +if (x_402 == 0) { -lean_object* x_401; -lean_dec(x_38); -x_401 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_401; +lean_object* x_403; +lean_dec(x_40); +x_403 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_403; } else { -lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; -x_402 = lean_unsigned_to_nat(0u); -x_403 = lean_array_fget(x_38, x_402); -lean_dec(x_38); -x_404 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__52; -x_405 = l_Lean_Expr_app___override(x_404, x_403); -x_406 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_405, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_406; +lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; +x_404 = lean_unsigned_to_nat(0u); +x_405 = lean_array_fget(x_40, x_404); +lean_dec(x_40); +x_406 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__52; +x_407 = l_Lean_Expr_app___override(x_406, x_405); +x_408 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_407, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_408; } } } } else { -lean_object* x_407; +lean_object* x_409; +lean_dec(x_39); +lean_dec(x_38); lean_dec(x_37); lean_dec(x_36); -lean_dec(x_35); -lean_dec(x_34); lean_dec(x_2); -x_407 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_407; +x_409 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_409; } } else { -lean_object* x_408; +lean_object* x_410; +lean_dec(x_38); +lean_dec(x_37); lean_dec(x_36); -lean_dec(x_35); -lean_dec(x_34); lean_dec(x_2); -x_408 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_408; +x_410 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_410; } } else { -lean_object* x_409; -lean_dec(x_35); -lean_dec(x_34); +lean_object* x_411; +lean_dec(x_37); +lean_dec(x_36); lean_dec(x_2); -x_409 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_409; +x_411 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_411; } } } @@ -14509,411 +14621,415 @@ x_3 = l_Lean_Expr_const___override(x_2, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleFinVal(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_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_Tactic_Omega_asLinearComboImpl_handleFinVal(lean_object* x_1, 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) { _start: { if (lean_obj_tag(x_4) == 1) { -lean_object* x_13; lean_object* x_14; -x_13 = lean_ctor_get(x_4, 0); -lean_inc(x_13); +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_4, 0); +lean_inc(x_15); lean_dec(x_4); -lean_inc(x_8); -x_14 = l_Lean_FVarId_getValue_x3f(x_13, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_14) == 0) +lean_inc(x_10); +x_16 = l_Lean_FVarId_getValue_x3f(x_15, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_15; -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -if (lean_obj_tag(x_15) == 0) +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +if (lean_obj_tag(x_17) == 0) { -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_ctor_get(x_14, 1); -lean_inc(x_16); -lean_dec(x_14); -x_17 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_16); -return x_17; +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_18); +return x_19; } 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; -x_18 = lean_ctor_get(x_14, 1); -lean_inc(x_18); -lean_dec(x_14); -x_19 = lean_ctor_get(x_15, 0); -lean_inc(x_19); -lean_dec(x_15); -x_20 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleFinVal___closed__14; -x_21 = l_Lean_mkAppB(x_20, x_3, x_19); -x_22 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__54; -x_23 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__11; -x_24 = l_Lean_mkApp3(x_22, x_23, x_2, x_21); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_20 = lean_ctor_get(x_16, 1); +lean_inc(x_20); +lean_dec(x_16); +x_21 = lean_ctor_get(x_17, 0); +lean_inc(x_21); +lean_dec(x_17); +x_22 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleFinVal___closed__14; +x_23 = l_Lean_mkAppB(x_22, x_3, x_21); +x_24 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__54; +x_25 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__11; +x_26 = l_Lean_mkApp3(x_24, x_25, x_2, x_23); +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_1); -x_25 = l_Lean_Elab_Tactic_Omega_mkEqReflWithExpectedType(x_1, x_24, x_8, x_9, x_10, x_11, x_18); -if (lean_obj_tag(x_25) == 0) +x_27 = l_Lean_Elab_Tactic_Omega_mkEqReflWithExpectedType(x_1, x_26, x_10, x_11, x_12, x_13, x_20); +if (lean_obj_tag(x_27) == 0) { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -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_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_26, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_27); -return x_28; +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_27, 1); +lean_inc(x_29); +lean_dec(x_27); +x_30 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_28, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_29); +return x_30; } else { -uint8_t x_29; +uint8_t x_31; +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_29 = !lean_is_exclusive(x_25); -if (x_29 == 0) +x_31 = !lean_is_exclusive(x_27); +if (x_31 == 0) { -return x_25; +return x_27; } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_25, 0); -x_31 = lean_ctor_get(x_25, 1); -lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_25); -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_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_27, 0); +x_33 = lean_ctor_get(x_27, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_27); +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_33; +uint8_t x_35; +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_33 = !lean_is_exclusive(x_14); -if (x_33 == 0) +x_35 = !lean_is_exclusive(x_16); +if (x_35 == 0) { -return x_14; +return x_16; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_14, 0); -x_35 = lean_ctor_get(x_14, 1); -lean_inc(x_35); -lean_inc(x_34); -lean_dec(x_14); -x_36 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -return x_36; +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_16, 0); +x_37 = lean_ctor_get(x_16, 1); +lean_inc(x_37); +lean_inc(x_36); +lean_dec(x_16); +x_38 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set(x_38, 1, x_37); +return x_38; } } } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_dec(x_2); -x_37 = l_Lean_Expr_getAppFnArgs(x_4); +x_39 = l_Lean_Expr_getAppFnArgs(x_4); lean_inc(x_3); -x_38 = l_Lean_Expr_nat_x3f(x_3); -x_39 = lean_ctor_get(x_37, 0); -lean_inc(x_39); -if (lean_obj_tag(x_39) == 1) -{ -lean_object* x_40; -x_40 = lean_ctor_get(x_39, 0); -lean_inc(x_40); -if (lean_obj_tag(x_40) == 1) -{ -lean_object* x_41; -x_41 = lean_ctor_get(x_40, 0); +x_40 = l_Lean_Expr_nat_x3f(x_3); +x_41 = lean_ctor_get(x_39, 0); lean_inc(x_41); -if (lean_obj_tag(x_41) == 0) +if (lean_obj_tag(x_41) == 1) { -lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; -x_42 = lean_ctor_get(x_37, 1); +lean_object* x_42; +x_42 = lean_ctor_get(x_41, 0); lean_inc(x_42); -lean_dec(x_37); -x_43 = lean_ctor_get(x_39, 1); +if (lean_obj_tag(x_42) == 1) +{ +lean_object* x_43; +x_43 = lean_ctor_get(x_42, 0); lean_inc(x_43); -lean_dec(x_39); -x_44 = lean_ctor_get(x_40, 1); -lean_inc(x_44); -lean_dec(x_40); -x_45 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__2; -x_46 = lean_string_dec_eq(x_44, x_45); -if (x_46 == 0) +if (lean_obj_tag(x_43) == 0) { -lean_object* x_47; uint8_t x_48; -x_47 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__2; -x_48 = lean_string_dec_eq(x_44, x_47); +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_39, 1); +lean_inc(x_44); +lean_dec(x_39); +x_45 = lean_ctor_get(x_41, 1); +lean_inc(x_45); +lean_dec(x_41); +x_46 = lean_ctor_get(x_42, 1); +lean_inc(x_46); +lean_dec(x_42); +x_47 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__2; +x_48 = lean_string_dec_eq(x_46, x_47); if (x_48 == 0) { lean_object* x_49; uint8_t x_50; -x_49 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__1; -x_50 = lean_string_dec_eq(x_44, x_49); +x_49 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__2; +x_50 = lean_string_dec_eq(x_46, x_49); if (x_50 == 0) { lean_object* x_51; uint8_t x_52; -lean_dec(x_3); -x_51 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__1; -x_52 = lean_string_dec_eq(x_44, x_51); -lean_dec(x_44); +x_51 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__1; +x_52 = lean_string_dec_eq(x_46, x_51); if (x_52 == 0) { -lean_object* x_53; -lean_dec(x_43); -lean_dec(x_42); -lean_dec(x_38); -x_53 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_53; +lean_object* x_53; uint8_t x_54; +lean_dec(x_3); +x_53 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__1; +x_54 = lean_string_dec_eq(x_46, x_53); +lean_dec(x_46); +if (x_54 == 0) +{ +lean_object* x_55; +lean_dec(x_45); +lean_dec(x_44); +lean_dec(x_40); +x_55 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_55; } else { -lean_object* x_54; uint8_t x_55; -x_54 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__37; -x_55 = lean_string_dec_eq(x_43, x_54); -lean_dec(x_43); -if (x_55 == 0) +lean_object* x_56; uint8_t x_57; +x_56 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__37; +x_57 = lean_string_dec_eq(x_45, x_56); +lean_dec(x_45); +if (x_57 == 0) { -lean_object* x_56; -lean_dec(x_42); -lean_dec(x_38); -x_56 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_56; +lean_object* x_58; +lean_dec(x_44); +lean_dec(x_40); +x_58 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_58; } else { -lean_object* x_57; lean_object* x_58; uint8_t x_59; -x_57 = lean_array_get_size(x_42); -x_58 = lean_unsigned_to_nat(3u); -x_59 = lean_nat_dec_eq(x_57, x_58); -lean_dec(x_57); -if (x_59 == 0) +lean_object* x_59; lean_object* x_60; uint8_t x_61; +x_59 = lean_array_get_size(x_44); +x_60 = lean_unsigned_to_nat(3u); +x_61 = lean_nat_dec_eq(x_59, x_60); +lean_dec(x_59); +if (x_61 == 0) { -lean_object* x_60; -lean_dec(x_42); -lean_dec(x_38); -x_60 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_60; +lean_object* x_62; +lean_dec(x_44); +lean_dec(x_40); +x_62 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_62; } else { -if (lean_obj_tag(x_38) == 0) +if (lean_obj_tag(x_40) == 0) { -lean_object* x_61; -lean_dec(x_42); -x_61 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_61; +lean_object* x_63; +lean_dec(x_44); +x_63 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_63; } else { -lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; uint8_t x_66; -x_62 = lean_ctor_get(x_38, 0); -lean_inc(x_62); -lean_dec(x_38); -x_63 = lean_unsigned_to_nat(1u); -x_64 = lean_array_fget(x_42, x_63); -lean_dec(x_42); -x_65 = lean_unsigned_to_nat(0u); -x_66 = lean_nat_dec_eq(x_62, x_65); -if (x_66 == 0) +lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; uint8_t x_68; +x_64 = lean_ctor_get(x_40, 0); +lean_inc(x_64); +lean_dec(x_40); +x_65 = lean_unsigned_to_nat(1u); +x_66 = lean_array_fget(x_44, x_65); +lean_dec(x_44); +x_67 = lean_unsigned_to_nat(0u); +x_68 = lean_nat_dec_eq(x_64, x_67); +if (x_68 == 0) { -if (lean_obj_tag(x_64) == 9) +if (lean_obj_tag(x_66) == 9) { -lean_object* x_67; -x_67 = lean_ctor_get(x_64, 0); -lean_inc(x_67); -lean_dec(x_64); -if (lean_obj_tag(x_67) == 0) +lean_object* x_69; +x_69 = lean_ctor_get(x_66, 0); +lean_inc(x_69); +lean_dec(x_66); +if (lean_obj_tag(x_69) == 0) { -uint8_t x_68; -x_68 = !lean_is_exclusive(x_67); -if (x_68 == 0) +uint8_t x_70; +x_70 = !lean_is_exclusive(x_69); +if (x_70 == 0) { -lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; -x_69 = lean_ctor_get(x_67, 0); -x_70 = lean_nat_mod(x_69, x_62); -lean_inc(x_70); -x_71 = l_Lean_mkNatLit(x_70); +lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_71 = lean_ctor_get(x_69, 0); +x_72 = lean_nat_mod(x_71, x_64); +lean_inc(x_72); +x_73 = l_Lean_mkNatLit(x_72); +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_72 = l_Lean_Meta_mkEqRefl(x_71, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_72) == 0) +x_74 = l_Lean_Meta_mkEqRefl(x_73, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_74) == 0) { -lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_73 = lean_ctor_get(x_72, 0); -lean_inc(x_73); -x_74 = lean_ctor_get(x_72, 1); -lean_inc(x_74); -lean_dec(x_72); -x_75 = lean_nat_sub(x_62, x_63); -lean_dec(x_62); -x_76 = l_Lean_mkNatLit(x_75); -x_77 = l_Lean_mkNatLit(x_69); -lean_ctor_set(x_67, 0, x_70); -x_78 = l_Lean_Expr_lit___override(x_67); -x_79 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleFinVal___closed__3; -x_80 = l_Lean_mkApp4(x_79, x_76, x_77, x_78, x_73); -x_81 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_80, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_74); -return x_81; +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 = 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_nat_sub(x_64, x_65); +lean_dec(x_64); +x_78 = l_Lean_mkNatLit(x_77); +x_79 = l_Lean_mkNatLit(x_71); +lean_ctor_set(x_69, 0, x_72); +x_80 = l_Lean_Expr_lit___override(x_69); +x_81 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleFinVal___closed__3; +x_82 = l_Lean_mkApp4(x_81, x_78, x_79, x_80, x_75); +x_83 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_82, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_76); +return x_83; } else { -uint8_t x_82; -lean_dec(x_70); -lean_free_object(x_67); -lean_dec(x_69); -lean_dec(x_62); +uint8_t x_84; +lean_dec(x_72); +lean_free_object(x_69); +lean_dec(x_71); +lean_dec(x_64); +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_82 = !lean_is_exclusive(x_72); -if (x_82 == 0) +x_84 = !lean_is_exclusive(x_74); +if (x_84 == 0) { -return x_72; +return x_74; } else { -lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_83 = lean_ctor_get(x_72, 0); -x_84 = lean_ctor_get(x_72, 1); -lean_inc(x_84); -lean_inc(x_83); -lean_dec(x_72); -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; +lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_85 = lean_ctor_get(x_74, 0); +x_86 = lean_ctor_get(x_74, 1); +lean_inc(x_86); +lean_inc(x_85); +lean_dec(x_74); +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_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_86 = lean_ctor_get(x_67, 0); -lean_inc(x_86); -lean_dec(x_67); -x_87 = lean_nat_mod(x_86, x_62); -lean_inc(x_87); -x_88 = l_Lean_mkNatLit(x_87); +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_88 = lean_ctor_get(x_69, 0); +lean_inc(x_88); +lean_dec(x_69); +x_89 = lean_nat_mod(x_88, x_64); +lean_inc(x_89); +x_90 = l_Lean_mkNatLit(x_89); +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_89 = l_Lean_Meta_mkEqRefl(x_88, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_89) == 0) +x_91 = l_Lean_Meta_mkEqRefl(x_90, x_10, x_11, x_12, x_13, x_14); +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; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; -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); -x_92 = lean_nat_sub(x_62, x_63); -lean_dec(x_62); -x_93 = l_Lean_mkNatLit(x_92); -x_94 = l_Lean_mkNatLit(x_86); -x_95 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_95, 0, x_87); -x_96 = l_Lean_Expr_lit___override(x_95); -x_97 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleFinVal___closed__3; -x_98 = l_Lean_mkApp4(x_97, x_93, x_94, x_96, x_90); -x_99 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_98, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_91); -return x_99; +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_92 = lean_ctor_get(x_91, 0); +lean_inc(x_92); +x_93 = lean_ctor_get(x_91, 1); +lean_inc(x_93); +lean_dec(x_91); +x_94 = lean_nat_sub(x_64, x_65); +lean_dec(x_64); +x_95 = l_Lean_mkNatLit(x_94); +x_96 = l_Lean_mkNatLit(x_88); +x_97 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_97, 0, x_89); +x_98 = l_Lean_Expr_lit___override(x_97); +x_99 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleFinVal___closed__3; +x_100 = l_Lean_mkApp4(x_99, x_95, x_96, x_98, x_92); +x_101 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_100, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_93); +return x_101; } else { -lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; -lean_dec(x_87); -lean_dec(x_86); -lean_dec(x_62); +lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; +lean_dec(x_89); +lean_dec(x_88); +lean_dec(x_64); +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_100 = lean_ctor_get(x_89, 0); -lean_inc(x_100); -x_101 = lean_ctor_get(x_89, 1); -lean_inc(x_101); -if (lean_is_exclusive(x_89)) { - lean_ctor_release(x_89, 0); - lean_ctor_release(x_89, 1); - x_102 = x_89; +x_102 = lean_ctor_get(x_91, 0); +lean_inc(x_102); +x_103 = lean_ctor_get(x_91, 1); +lean_inc(x_103); +if (lean_is_exclusive(x_91)) { + lean_ctor_release(x_91, 0); + lean_ctor_release(x_91, 1); + x_104 = x_91; } else { - lean_dec_ref(x_89); - x_102 = lean_box(0); + lean_dec_ref(x_91); + x_104 = lean_box(0); } -if (lean_is_scalar(x_102)) { - x_103 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_104)) { + x_105 = lean_alloc_ctor(1, 2, 0); } else { - x_103 = x_102; + x_105 = x_104; } -lean_ctor_set(x_103, 0, x_100); -lean_ctor_set(x_103, 1, x_101); -return x_103; +lean_ctor_set(x_105, 0, x_102); +lean_ctor_set(x_105, 1, x_103); +return x_105; } } } else { -lean_object* x_104; -lean_dec(x_67); -lean_dec(x_62); -x_104 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_104; +lean_object* x_106; +lean_dec(x_69); +lean_dec(x_64); +x_106 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_106; } } else { -lean_object* x_105; +lean_object* x_107; +lean_dec(x_66); lean_dec(x_64); -lean_dec(x_62); -x_105 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_105; +x_107 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_107; } } else { -lean_object* x_106; +lean_object* x_108; +lean_dec(x_66); lean_dec(x_64); -lean_dec(x_62); -x_106 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_106; +x_108 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_108; } } } @@ -14922,59 +15038,59 @@ return x_106; } else { -lean_object* x_107; uint8_t x_108; -lean_dec(x_44); -x_107 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__47; -x_108 = lean_string_dec_eq(x_43, x_107); -lean_dec(x_43); -if (x_108 == 0) +lean_object* x_109; uint8_t x_110; +lean_dec(x_46); +x_109 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__47; +x_110 = lean_string_dec_eq(x_45, x_109); +lean_dec(x_45); +if (x_110 == 0) { -lean_object* x_109; -lean_dec(x_42); -lean_dec(x_38); +lean_object* x_111; +lean_dec(x_44); +lean_dec(x_40); lean_dec(x_3); -x_109 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_109; +x_111 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_111; } else { -lean_object* x_110; lean_object* x_111; uint8_t x_112; -x_110 = lean_array_get_size(x_42); -x_111 = lean_unsigned_to_nat(6u); -x_112 = lean_nat_dec_eq(x_110, x_111); -lean_dec(x_110); -if (x_112 == 0) +lean_object* x_112; lean_object* x_113; uint8_t x_114; +x_112 = lean_array_get_size(x_44); +x_113 = lean_unsigned_to_nat(6u); +x_114 = lean_nat_dec_eq(x_112, x_113); +lean_dec(x_112); +if (x_114 == 0) { -lean_object* x_113; -lean_dec(x_42); -lean_dec(x_38); +lean_object* x_115; +lean_dec(x_44); +lean_dec(x_40); lean_dec(x_3); -x_113 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_113; +x_115 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_115; } else { -if (lean_obj_tag(x_38) == 0) +if (lean_obj_tag(x_40) == 0) { -lean_object* x_114; -lean_dec(x_42); +lean_object* x_116; +lean_dec(x_44); lean_dec(x_3); -x_114 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_114; +x_116 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_116; } 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_dec(x_38); -x_115 = lean_unsigned_to_nat(4u); -x_116 = lean_array_fget(x_42, x_115); -x_117 = lean_unsigned_to_nat(5u); -x_118 = lean_array_fget(x_42, x_117); -lean_dec(x_42); -x_119 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleFinVal___closed__6; -x_120 = l_Lean_mkApp3(x_119, x_3, x_116, x_118); -x_121 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_120, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_121; +lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; +lean_dec(x_40); +x_117 = lean_unsigned_to_nat(4u); +x_118 = lean_array_fget(x_44, x_117); +x_119 = lean_unsigned_to_nat(5u); +x_120 = lean_array_fget(x_44, x_119); +lean_dec(x_44); +x_121 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleFinVal___closed__6; +x_122 = l_Lean_mkApp3(x_121, x_3, x_118, x_120); +x_123 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_122, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_123; } } } @@ -14982,145 +15098,178 @@ return x_121; } else { -lean_object* x_122; uint8_t x_123; -lean_dec(x_44); -lean_dec(x_38); -x_122 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__45; -x_123 = lean_string_dec_eq(x_43, x_122); -lean_dec(x_43); -if (x_123 == 0) +lean_object* x_124; uint8_t x_125; +lean_dec(x_46); +lean_dec(x_40); +x_124 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__45; +x_125 = lean_string_dec_eq(x_45, x_124); +lean_dec(x_45); +if (x_125 == 0) { -lean_object* x_124; -lean_dec(x_42); +lean_object* x_126; +lean_dec(x_44); lean_dec(x_3); -x_124 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_124; +x_126 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_126; } else { -lean_object* x_125; lean_object* x_126; uint8_t x_127; -x_125 = lean_array_get_size(x_42); -x_126 = lean_unsigned_to_nat(6u); -x_127 = lean_nat_dec_eq(x_125, x_126); -lean_dec(x_125); -if (x_127 == 0) +lean_object* x_127; lean_object* x_128; uint8_t x_129; +x_127 = lean_array_get_size(x_44); +x_128 = lean_unsigned_to_nat(6u); +x_129 = lean_nat_dec_eq(x_127, x_128); +lean_dec(x_127); +if (x_129 == 0) { -lean_object* x_128; -lean_dec(x_42); +lean_object* x_130; +lean_dec(x_44); lean_dec(x_3); -x_128 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_128; +x_130 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_130; } else { -lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; -x_129 = lean_unsigned_to_nat(4u); -x_130 = lean_array_fget(x_42, x_129); -x_131 = lean_unsigned_to_nat(5u); -x_132 = lean_array_fget(x_42, x_131); -lean_dec(x_42); -x_133 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleFinVal___closed__9; -x_134 = l_Lean_mkApp3(x_133, x_3, x_130, x_132); -x_135 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_134, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_135; +lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; +x_131 = lean_unsigned_to_nat(4u); +x_132 = lean_array_fget(x_44, x_131); +x_133 = lean_unsigned_to_nat(5u); +x_134 = lean_array_fget(x_44, x_133); +lean_dec(x_44); +x_135 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleFinVal___closed__9; +x_136 = l_Lean_mkApp3(x_135, x_3, x_132, x_134); +x_137 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_136, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_137; } } } } else { -lean_object* x_136; uint8_t x_137; -lean_dec(x_44); -lean_dec(x_38); -x_136 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__3; -x_137 = lean_string_dec_eq(x_43, x_136); -lean_dec(x_43); -if (x_137 == 0) +lean_object* x_138; uint8_t x_139; +lean_dec(x_46); +lean_dec(x_40); +x_138 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__3; +x_139 = lean_string_dec_eq(x_45, x_138); +lean_dec(x_45); +if (x_139 == 0) { -lean_object* x_138; -lean_dec(x_42); +lean_object* x_140; +lean_dec(x_44); lean_dec(x_3); -x_138 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_138; +x_140 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_140; } else { -lean_object* x_139; lean_object* x_140; uint8_t x_141; -x_139 = lean_array_get_size(x_42); -x_140 = lean_unsigned_to_nat(6u); -x_141 = lean_nat_dec_eq(x_139, x_140); -lean_dec(x_139); -if (x_141 == 0) +lean_object* x_141; lean_object* x_142; uint8_t x_143; +x_141 = lean_array_get_size(x_44); +x_142 = lean_unsigned_to_nat(6u); +x_143 = lean_nat_dec_eq(x_141, x_142); +lean_dec(x_141); +if (x_143 == 0) { -lean_object* x_142; -lean_dec(x_42); +lean_object* x_144; +lean_dec(x_44); lean_dec(x_3); -x_142 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_142; +x_144 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_144; } else { -lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; -x_143 = lean_unsigned_to_nat(4u); -x_144 = lean_array_fget(x_42, x_143); -x_145 = lean_unsigned_to_nat(5u); -x_146 = lean_array_fget(x_42, x_145); -lean_dec(x_42); -x_147 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleFinVal___closed__12; -x_148 = l_Lean_mkApp3(x_147, x_3, x_144, x_146); -x_149 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_148, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_149; +lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; +x_145 = lean_unsigned_to_nat(4u); +x_146 = lean_array_fget(x_44, x_145); +x_147 = lean_unsigned_to_nat(5u); +x_148 = lean_array_fget(x_44, x_147); +lean_dec(x_44); +x_149 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleFinVal___closed__12; +x_150 = l_Lean_mkApp3(x_149, x_3, x_146, x_148); +x_151 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_150, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_151; } } } } else { -lean_object* x_150; +lean_object* x_152; +lean_dec(x_43); +lean_dec(x_42); lean_dec(x_41); lean_dec(x_40); lean_dec(x_39); -lean_dec(x_38); -lean_dec(x_37); lean_dec(x_3); -x_150 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_150; +x_152 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_152; } } else { -lean_object* x_151; +lean_object* x_153; +lean_dec(x_42); +lean_dec(x_41); lean_dec(x_40); lean_dec(x_39); -lean_dec(x_38); -lean_dec(x_37); lean_dec(x_3); -x_151 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_151; +x_153 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_153; } } else { -lean_object* x_152; +lean_object* x_154; +lean_dec(x_41); +lean_dec(x_40); lean_dec(x_39); -lean_dec(x_38); -lean_dec(x_37); lean_dec(x_3); -x_152 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_152; +x_154 = l_Lean_Elab_Tactic_Omega_mkAtomLinearCombo(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_154; } } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___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_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___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_11; -x_11 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +uint8_t x_12; lean_object* x_13; +x_12 = lean_unbox(x_5); +lean_dec(x_5); +x_13 = l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1(x_1, x_2, x_3, x_4, x_12, x_6, x_7, x_8, x_9, x_10, x_11); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; lean_object* x_14; +x_13 = lean_unbox(x_6); +lean_dec(x_6); +x_14 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___lambda__1(x_1, x_2, x_3, x_4, x_5, x_13, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_7); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -return x_11; +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___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_6); +lean_dec(x_6); +x_14 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___lambda__2(x_1, x_2, x_3, x_4, x_5, x_13, x_7, x_8, x_9, x_10, x_11, x_12); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, 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_6); +lean_dec(x_6); +x_14 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite(x_1, x_2, x_3, x_4, x_5, x_13, x_7, x_8, x_9, x_10, x_11, x_12); +return x_14; } } LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__2___boxed(lean_object* x_1, lean_object* x_2) { @@ -15133,6 +15282,16 @@ lean_dec(x_1); return x_3; } } +LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__3___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, 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_5); +lean_dec(x_5); +x_13 = l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__3___rarg(x_1, x_2, x_3, x_4, x_12, x_6, x_7, x_8, x_9, x_10, x_11); +return x_13; +} +} LEAN_EXPORT lean_object* l_Lean_AssocList_contains___at_Lean_Elab_Tactic_Omega_asLinearCombo___spec__5___boxed(lean_object* x_1, lean_object* x_2) { _start: { @@ -15144,48 +15303,77 @@ x_4 = lean_box(x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearCombo___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_Elab_Tactic_Omega_asLinearCombo___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_8; -x_8 = l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +uint8_t x_10; lean_object* x_11; +x_10 = lean_unbox(x_3); +lean_dec(x_3); +x_11 = l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__1(x_1, x_2, x_10, 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_8; +return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearCombo___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_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearCombo___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_9; -x_9 = l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +uint8_t x_12; lean_object* x_13; +x_12 = lean_unbox(x_5); +lean_dec(x_5); +x_13 = l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__2(x_1, x_2, x_3, x_4, x_12, x_6, x_7, x_8, x_9, x_10, x_11); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +uint8_t x_11; lean_object* x_12; +x_11 = lean_unbox(x_4); +lean_dec(x_4); +x_12 = l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__3(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_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; +return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearCombo___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_Elab_Tactic_Omega_asLinearCombo___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) { _start: { -lean_object* x_12; -x_12 = l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +uint8_t x_14; lean_object* x_15; +x_14 = lean_unbox(x_7); +lean_dec(x_7); +x_15 = l_Lean_Elab_Tactic_Omega_asLinearCombo___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_14, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -return x_12; +return x_15; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearCombo___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_5); +lean_dec(x_5); +x_13 = l_Lean_Elab_Tactic_Omega_asLinearCombo(x_1, x_2, x_3, x_4, x_12, x_6, x_7, x_8, x_9, x_10, x_11); +return x_13; } } LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Omega_asLinearComboImpl___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { @@ -15211,13 +15399,16 @@ x_3 = lean_box(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___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_8; -x_8 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +uint8_t x_10; lean_object* x_11; +x_10 = lean_unbox(x_3); +lean_dec(x_3); +x_11 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__2(x_1, x_2, x_10, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_4); lean_dec(x_2); -return x_8; +return x_11; } } LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__3___boxed(lean_object** _args) { @@ -15238,44 +15429,104 @@ 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]; +lean_object* x_19 = _args[18]; _start: { -lean_object* x_18; -x_18 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___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, x_15, x_16, x_17); +uint8_t x_20; lean_object* x_21; +x_20 = lean_unbox(x_13); +lean_dec(x_13); +x_21 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___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_20, x_14, x_15, x_16, x_17, x_18, x_19); lean_dec(x_3); lean_dec(x_2); -return x_18; +return x_21; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___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_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___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_12; -x_12 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +uint8_t x_13; lean_object* x_14; +x_13 = lean_unbox(x_6); +lean_dec(x_6); +x_14 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__4(x_1, x_2, x_3, x_4, x_5, x_13, x_7, x_8, x_9, x_10, x_11, x_12); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___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) { +_start: +{ +uint8_t x_14; lean_object* x_15; +x_14 = lean_unbox(x_7); +lean_dec(x_7); +x_15 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_14, x_8, x_9, x_10, x_11, x_12, x_13); lean_dec(x_1); -return x_12; +return x_15; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___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_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___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: { -lean_object* x_14; -x_14 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___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); +uint8_t x_16; lean_object* x_17; +x_16 = lean_unbox(x_9); +lean_dec(x_9); +x_17 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___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_13, x_14, x_15); lean_dec(x_2); lean_dec(x_1); -return x_14; +return x_17; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___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_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___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: { -lean_object* x_14; -x_14 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___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_13); +uint8_t x_16; lean_object* x_17; +x_16 = lean_unbox(x_9); +lean_dec(x_9); +x_17 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___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_13, x_14, x_15); lean_dec(x_2); lean_dec(x_1); +return x_17; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___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) { +_start: +{ +uint8_t x_13; lean_object* x_14; +x_13 = lean_unbox(x_6); +lean_dec(x_6); +x_14 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8(x_1, x_2, x_3, x_4, x_5, x_13, x_7, x_8, x_9, x_10, x_11, x_12); return x_14; } } +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_5); +lean_dec(x_5); +x_13 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl(x_1, x_2, x_3, x_4, x_12, x_6, x_7, x_8, x_9, x_10, x_11); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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_7); +lean_dec(x_7); +x_15 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast(x_1, x_2, x_3, x_4, x_5, x_6, x_14, x_8, x_9, x_10, x_11, x_12, x_13); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleFinVal___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, 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_8); +lean_dec(x_8); +x_16 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleFinVal(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_15, x_9, x_10, x_11, x_12, x_13, x_14); +return x_16; +} +} static lean_object* _init_l_Lean_Elab_Tactic_Omega_MetaProblem_trivial___closed__1() { _start: { @@ -15368,226 +15619,226 @@ return x_5; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___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_Elab_Tactic_Omega_MetaProblem_addIntEquality___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_object* x_11, lean_object* x_12) { _start: { -lean_object* x_11; +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); -x_11 = l_Lean_Meta_mkEqSymm(x_2, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_11) == 0) +x_13 = l_Lean_Meta_mkEqSymm(x_2, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 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_Meta_mkEqTrans(x_12, x_1, x_6, x_7, x_8, x_9, x_13); -return x_14; +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_Meta_mkEqTrans(x_14, x_1, x_8, x_9, x_10, x_11, x_15); +return x_16; } else { -uint8_t x_15; +uint8_t 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_1); -x_15 = !lean_is_exclusive(x_11); -if (x_15 == 0) +x_17 = !lean_is_exclusive(x_13); +if (x_17 == 0) { -return x_11; +return x_13; } else { -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; +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_13, 0); +x_19 = lean_ctor_get(x_13, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_13); +x_20 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 1, x_19); +return x_20; } } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___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_Elab_Tactic_Omega_MetaProblem_addIntEquality___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, 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) { _start: { -uint8_t x_15; -x_15 = !lean_is_exclusive(x_1); -if (x_15 == 0) +uint8_t x_17; +x_17 = !lean_is_exclusive(x_1); +if (x_17 == 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; -x_16 = lean_ctor_get(x_1, 0); -x_17 = lean_ctor_get(x_1, 1); -x_18 = lean_ctor_get(x_1, 2); -x_19 = lean_ctor_get(x_1, 3); -x_20 = lean_ctor_get(x_2, 0); -lean_inc(x_20); -x_21 = lean_ctor_get(x_2, 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; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_18 = lean_ctor_get(x_1, 0); +x_19 = lean_ctor_get(x_1, 1); +x_20 = lean_ctor_get(x_1, 2); +x_21 = lean_ctor_get(x_1, 3); +x_22 = lean_ctor_get(x_2, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_2, 1); +lean_inc(x_23); lean_dec(x_2); -x_22 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___lambda__1___boxed), 10, 1); -lean_closure_set(x_22, 0, x_3); -x_23 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg), 10, 2); -lean_closure_set(x_23, 0, x_4); -lean_closure_set(x_23, 1, x_22); -x_24 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_24, 0, x_23); -x_25 = l_Lean_Elab_Tactic_Omega_Problem_addEquality(x_16, x_20, x_21, x_24); -x_26 = l_Lean_Elab_Tactic_Omega_Problem_solveEqualities(x_25, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_26) == 0) +x_24 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___lambda__1___boxed), 12, 1); +lean_closure_set(x_24, 0, x_3); +x_25 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg___boxed), 12, 2); +lean_closure_set(x_25, 0, x_4); +lean_closure_set(x_25, 1, x_24); +x_26 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_26, 0, x_25); +x_27 = l_Lean_Elab_Tactic_Omega_Problem_addEquality(x_18, x_22, x_23, x_26); +x_28 = l_Lean_Elab_Tactic_Omega_Problem_solveEqualities(x_27, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_28) == 0) { -uint8_t x_27; -x_27 = !lean_is_exclusive(x_26); -if (x_27 == 0) +uint8_t x_29; +x_29 = !lean_is_exclusive(x_28); +if (x_29 == 0) { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_26, 0); -x_29 = l_Lean_HashSet_toList___at_Lean_Elab_Tactic_Omega_lookup___spec__10(x_5); -x_30 = l_List_appendTR___rarg(x_29, x_17); -lean_ctor_set(x_1, 1, x_30); -lean_ctor_set(x_1, 0, x_28); -lean_ctor_set(x_26, 0, x_1); -return x_26; +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_28, 0); +x_31 = l_Lean_HashSet_toList___at_Lean_Elab_Tactic_Omega_lookup___spec__10(x_5); +x_32 = l_List_appendTR___rarg(x_31, x_19); +lean_ctor_set(x_1, 1, x_32); +lean_ctor_set(x_1, 0, x_30); +lean_ctor_set(x_28, 0, x_1); +return x_28; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_31 = lean_ctor_get(x_26, 0); -x_32 = lean_ctor_get(x_26, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_26); -x_33 = l_Lean_HashSet_toList___at_Lean_Elab_Tactic_Omega_lookup___spec__10(x_5); -x_34 = l_List_appendTR___rarg(x_33, x_17); -lean_ctor_set(x_1, 1, x_34); -lean_ctor_set(x_1, 0, x_31); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_1); -lean_ctor_set(x_35, 1, x_32); -return x_35; +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_33 = lean_ctor_get(x_28, 0); +x_34 = lean_ctor_get(x_28, 1); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_28); +x_35 = l_Lean_HashSet_toList___at_Lean_Elab_Tactic_Omega_lookup___spec__10(x_5); +x_36 = l_List_appendTR___rarg(x_35, x_19); +lean_ctor_set(x_1, 1, x_36); +lean_ctor_set(x_1, 0, x_33); +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_1); +lean_ctor_set(x_37, 1, x_34); +return x_37; } } else { -uint8_t x_36; +uint8_t x_38; lean_free_object(x_1); +lean_dec(x_21); +lean_dec(x_20); lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_17); lean_dec(x_5); -x_36 = !lean_is_exclusive(x_26); -if (x_36 == 0) +x_38 = !lean_is_exclusive(x_28); +if (x_38 == 0) { -return x_26; +return x_28; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_26, 0); -x_38 = lean_ctor_get(x_26, 1); -lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_26); -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_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_28, 0); +x_40 = lean_ctor_get(x_28, 1); +lean_inc(x_40); +lean_inc(x_39); +lean_dec(x_28); +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_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; 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_40 = lean_ctor_get(x_1, 0); -x_41 = lean_ctor_get(x_1, 1); -x_42 = lean_ctor_get(x_1, 2); -x_43 = lean_ctor_get(x_1, 3); +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; 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_42 = lean_ctor_get(x_1, 0); +x_43 = lean_ctor_get(x_1, 1); +x_44 = lean_ctor_get(x_1, 2); +x_45 = lean_ctor_get(x_1, 3); +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_1); -x_44 = lean_ctor_get(x_2, 0); -lean_inc(x_44); -x_45 = lean_ctor_get(x_2, 1); -lean_inc(x_45); +x_46 = lean_ctor_get(x_2, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_2, 1); +lean_inc(x_47); lean_dec(x_2); -x_46 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___lambda__1___boxed), 10, 1); -lean_closure_set(x_46, 0, x_3); -x_47 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg), 10, 2); -lean_closure_set(x_47, 0, x_4); -lean_closure_set(x_47, 1, x_46); -x_48 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_48, 0, x_47); -x_49 = l_Lean_Elab_Tactic_Omega_Problem_addEquality(x_40, x_44, x_45, x_48); -x_50 = l_Lean_Elab_Tactic_Omega_Problem_solveEqualities(x_49, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -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; -x_51 = lean_ctor_get(x_50, 0); -lean_inc(x_51); -x_52 = lean_ctor_get(x_50, 1); -lean_inc(x_52); -if (lean_is_exclusive(x_50)) { - lean_ctor_release(x_50, 0); - lean_ctor_release(x_50, 1); - x_53 = x_50; +x_48 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___lambda__1___boxed), 12, 1); +lean_closure_set(x_48, 0, x_3); +x_49 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg___boxed), 12, 2); +lean_closure_set(x_49, 0, x_4); +lean_closure_set(x_49, 1, x_48); +x_50 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_50, 0, x_49); +x_51 = l_Lean_Elab_Tactic_Omega_Problem_addEquality(x_42, x_46, x_47, x_50); +x_52 = l_Lean_Elab_Tactic_Omega_Problem_solveEqualities(x_51, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_52) == 0) +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_53 = lean_ctor_get(x_52, 0); +lean_inc(x_53); +x_54 = lean_ctor_get(x_52, 1); +lean_inc(x_54); +if (lean_is_exclusive(x_52)) { + lean_ctor_release(x_52, 0); + lean_ctor_release(x_52, 1); + x_55 = x_52; } else { - lean_dec_ref(x_50); - x_53 = lean_box(0); + lean_dec_ref(x_52); + x_55 = lean_box(0); } -x_54 = l_Lean_HashSet_toList___at_Lean_Elab_Tactic_Omega_lookup___spec__10(x_5); -x_55 = l_List_appendTR___rarg(x_54, x_41); -x_56 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_56, 0, x_51); -lean_ctor_set(x_56, 1, x_55); -lean_ctor_set(x_56, 2, x_42); -lean_ctor_set(x_56, 3, x_43); -if (lean_is_scalar(x_53)) { - x_57 = lean_alloc_ctor(0, 2, 0); +x_56 = l_Lean_HashSet_toList___at_Lean_Elab_Tactic_Omega_lookup___spec__10(x_5); +x_57 = l_List_appendTR___rarg(x_56, x_43); +x_58 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_58, 0, x_53); +lean_ctor_set(x_58, 1, x_57); +lean_ctor_set(x_58, 2, x_44); +lean_ctor_set(x_58, 3, x_45); +if (lean_is_scalar(x_55)) { + x_59 = lean_alloc_ctor(0, 2, 0); } else { - x_57 = x_53; + x_59 = x_55; } -lean_ctor_set(x_57, 0, x_56); -lean_ctor_set(x_57, 1, x_52); -return x_57; +lean_ctor_set(x_59, 0, x_58); +lean_ctor_set(x_59, 1, x_54); +return x_59; } else { -lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; +lean_dec(x_45); +lean_dec(x_44); lean_dec(x_43); -lean_dec(x_42); -lean_dec(x_41); lean_dec(x_5); -x_58 = lean_ctor_get(x_50, 0); -lean_inc(x_58); -x_59 = lean_ctor_get(x_50, 1); -lean_inc(x_59); -if (lean_is_exclusive(x_50)) { - lean_ctor_release(x_50, 0); - lean_ctor_release(x_50, 1); - x_60 = x_50; +x_60 = lean_ctor_get(x_52, 0); +lean_inc(x_60); +x_61 = lean_ctor_get(x_52, 1); +lean_inc(x_61); +if (lean_is_exclusive(x_52)) { + lean_ctor_release(x_52, 0); + lean_ctor_release(x_52, 1); + x_62 = x_52; } else { - lean_dec_ref(x_50); - x_60 = lean_box(0); + lean_dec_ref(x_52); + x_62 = lean_box(0); } -if (lean_is_scalar(x_60)) { - x_61 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_62)) { + x_63 = lean_alloc_ctor(1, 2, 0); } else { - x_61 = x_60; + x_63 = x_62; } -lean_ctor_set(x_61, 0, x_58); -lean_ctor_set(x_61, 1, x_59); -return x_61; +lean_ctor_set(x_63, 0, x_60); +lean_ctor_set(x_63, 1, x_61); +return x_63; } } } @@ -15634,224 +15885,226 @@ x_1 = lean_mk_string_from_bytes("-", 1); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_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_Omega_MetaProblem_addIntEquality(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_object* x_12, lean_object* x_13) { _start: { -lean_object* x_12; +lean_object* 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_6); lean_inc(x_5); lean_inc(x_4); -x_12 = l_Lean_Elab_Tactic_Omega_asLinearCombo(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_12) == 0) +x_14 = l_Lean_Elab_Tactic_Omega_asLinearCombo(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_14) == 0) { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_76; lean_object* x_77; lean_object* x_78; uint8_t x_79; -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_13, 1); -lean_inc(x_14); -x_15 = lean_ctor_get(x_12, 1); +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_78; lean_object* x_79; lean_object* x_80; uint8_t x_81; +x_15 = lean_ctor_get(x_14, 0); lean_inc(x_15); -lean_dec(x_12); -x_16 = lean_ctor_get(x_13, 0); +x_16 = lean_ctor_get(x_15, 1); lean_inc(x_16); -lean_dec(x_13); -x_17 = lean_ctor_get(x_14, 0); +x_17 = lean_ctor_get(x_14, 1); lean_inc(x_17); -x_18 = lean_ctor_get(x_14, 1); -lean_inc(x_18); lean_dec(x_14); -x_19 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__2; -x_76 = lean_ctor_get(x_18, 1); -lean_inc(x_76); -lean_dec(x_18); -x_77 = lean_array_get_size(x_76); -x_78 = lean_unsigned_to_nat(0u); -x_79 = lean_nat_dec_lt(x_78, x_77); -if (x_79 == 0) +x_18 = lean_ctor_get(x_15, 0); +lean_inc(x_18); +lean_dec(x_15); +x_19 = lean_ctor_get(x_16, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_16, 1); +lean_inc(x_20); +lean_dec(x_16); +x_21 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__2; +x_78 = lean_ctor_get(x_20, 1); +lean_inc(x_78); +lean_dec(x_20); +x_79 = lean_array_get_size(x_78); +x_80 = lean_unsigned_to_nat(0u); +x_81 = lean_nat_dec_lt(x_80, x_79); +if (x_81 == 0) { -lean_object* x_80; -lean_dec(x_77); -lean_dec(x_76); -x_80 = l_Lean_Elab_Tactic_Omega_MetaProblem_problem___default___closed__2; -x_20 = x_80; -goto block_75; +lean_object* x_82; +lean_dec(x_79); +lean_dec(x_78); +x_82 = l_Lean_Elab_Tactic_Omega_MetaProblem_problem___default___closed__2; +x_22 = x_82; +goto block_77; } else { -uint8_t x_81; -x_81 = lean_nat_dec_le(x_77, x_77); -if (x_81 == 0) +uint8_t x_83; +x_83 = lean_nat_dec_le(x_79, x_79); +if (x_83 == 0) { -lean_object* x_82; -lean_dec(x_77); -lean_dec(x_76); -x_82 = l_Lean_Elab_Tactic_Omega_MetaProblem_problem___default___closed__2; -x_20 = x_82; -goto block_75; +lean_object* x_84; +lean_dec(x_79); +lean_dec(x_78); +x_84 = l_Lean_Elab_Tactic_Omega_MetaProblem_problem___default___closed__2; +x_22 = x_84; +goto block_77; } else { -size_t x_83; size_t x_84; lean_object* x_85; lean_object* x_86; -x_83 = 0; -x_84 = lean_usize_of_nat(x_77); -lean_dec(x_77); -x_85 = l_Lean_Elab_Tactic_Omega_MetaProblem_problem___default___closed__2; +size_t x_85; size_t x_86; lean_object* x_87; lean_object* x_88; +x_85 = 0; +x_86 = lean_usize_of_nat(x_79); +lean_dec(x_79); +x_87 = l_Lean_Elab_Tactic_Omega_MetaProblem_problem___default___closed__2; lean_inc(x_1); -x_86 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___spec__2(x_1, x_76, x_83, x_84, x_85); -lean_dec(x_76); -x_20 = x_86; -goto block_75; +x_88 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___spec__2(x_1, x_78, x_85, x_86, x_87); +lean_dec(x_78); +x_22 = x_88; +goto block_77; } } -block_75: -{ -lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_21 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_19, x_4, x_5, x_6, x_7, x_8, x_9, x_10, 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) +block_77: { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_21, 1); +lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_23 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_21, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_17); +x_24 = lean_ctor_get(x_23, 0); lean_inc(x_24); -lean_dec(x_21); -x_25 = lean_box(0); -x_26 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___lambda__2(x_1, x_16, x_2, x_17, x_20, x_25, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_24); -return x_26; +x_25 = lean_unbox(x_24); +lean_dec(x_24); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_23, 1); +lean_inc(x_26); +lean_dec(x_23); +x_27 = lean_box(0); +x_28 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___lambda__2(x_1, x_18, x_2, x_19, x_22, x_27, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_26); +return x_28; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_27 = lean_ctor_get(x_21, 1); -lean_inc(x_27); -lean_dec(x_21); -x_28 = lean_ctor_get(x_16, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_16, 1); +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; 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_29 = lean_ctor_get(x_23, 1); lean_inc(x_29); -x_30 = lean_unsigned_to_nat(0u); -x_31 = l_List_enumFrom___rarg(x_30, x_29); -x_32 = lean_box(0); -x_33 = l_List_mapTR_loop___at_Lean_Omega_LinearCombo_instToStringLinearCombo___spec__1(x_31, x_32); -x_34 = l_Lean_Elab_Tactic_Omega_MetaProblem_problem___default___lambda__1___closed__1; -x_35 = l_List_foldl___at_String_join___spec__1(x_34, x_33); -lean_dec(x_33); -x_36 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__8; -x_37 = lean_int_dec_lt(x_28, x_36); -if (x_37 == 0) -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; -x_38 = lean_nat_abs(x_28); -lean_dec(x_28); -x_39 = l_Nat_repr(x_38); -x_40 = lean_string_append(x_34, x_39); -lean_dec(x_39); -x_41 = lean_string_append(x_40, x_34); -x_42 = lean_string_append(x_41, x_35); +lean_dec(x_23); +x_30 = lean_ctor_get(x_18, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_18, 1); +lean_inc(x_31); +x_32 = lean_unsigned_to_nat(0u); +x_33 = l_List_enumFrom___rarg(x_32, x_31); +x_34 = lean_box(0); +x_35 = l_List_mapTR_loop___at_Lean_Omega_LinearCombo_instToStringLinearCombo___spec__1(x_33, x_34); +x_36 = l_Lean_Elab_Tactic_Omega_MetaProblem_problem___default___lambda__1___closed__1; +x_37 = l_List_foldl___at_String_join___spec__1(x_36, x_35); lean_dec(x_35); -x_43 = lean_string_append(x_42, x_34); -x_44 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_44, 0, x_43); -x_45 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_45, 0, x_44); -x_46 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___closed__2; -x_47 = lean_alloc_ctor(7, 2, 0); +x_38 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__8; +x_39 = lean_int_dec_lt(x_30, x_38); +if (x_39 == 0) +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_40 = lean_nat_abs(x_30); +lean_dec(x_30); +x_41 = l_Nat_repr(x_40); +x_42 = lean_string_append(x_36, x_41); +lean_dec(x_41); +x_43 = lean_string_append(x_42, x_36); +x_44 = lean_string_append(x_43, x_37); +lean_dec(x_37); +x_45 = lean_string_append(x_44, x_36); +x_46 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_46, 0, x_45); +x_47 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_45); -x_48 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___closed__4; +x_48 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___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 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_19, x_49, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_27); -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 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___lambda__2(x_1, x_16, x_2, x_17, x_20, x_51, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_52); -lean_dec(x_51); -return x_53; +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_47); +x_50 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___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 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_21, x_51, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_29); +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_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___lambda__2(x_1, x_18, x_2, x_19, x_22, x_53, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_54); +lean_dec(x_53); +return x_55; } else { -lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; -x_54 = lean_nat_abs(x_28); -lean_dec(x_28); -x_55 = lean_unsigned_to_nat(1u); -x_56 = lean_nat_sub(x_54, x_55); -lean_dec(x_54); -x_57 = lean_nat_add(x_56, x_55); +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_56 = lean_nat_abs(x_30); +lean_dec(x_30); +x_57 = lean_unsigned_to_nat(1u); +x_58 = lean_nat_sub(x_56, x_57); lean_dec(x_56); -x_58 = l_Nat_repr(x_57); -x_59 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___closed__5; -x_60 = lean_string_append(x_59, x_58); +x_59 = lean_nat_add(x_58, x_57); lean_dec(x_58); -x_61 = lean_string_append(x_34, x_60); +x_60 = l_Nat_repr(x_59); +x_61 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___closed__5; +x_62 = lean_string_append(x_61, x_60); lean_dec(x_60); -x_62 = lean_string_append(x_61, x_34); -x_63 = lean_string_append(x_62, x_35); -lean_dec(x_35); -x_64 = lean_string_append(x_63, x_34); -x_65 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_65, 0, x_64); -x_66 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_66, 0, x_65); -x_67 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___closed__2; -x_68 = lean_alloc_ctor(7, 2, 0); +x_63 = lean_string_append(x_36, x_62); +lean_dec(x_62); +x_64 = lean_string_append(x_63, x_36); +x_65 = lean_string_append(x_64, x_37); +lean_dec(x_37); +x_66 = lean_string_append(x_65, x_36); +x_67 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_67, 0, x_66); +x_68 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_68, 0, x_67); -lean_ctor_set(x_68, 1, x_66); -x_69 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___closed__4; +x_69 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___closed__2; x_70 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_70, 0, x_68); -lean_ctor_set(x_70, 1, x_69); -x_71 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_19, x_70, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_27); -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 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___lambda__2(x_1, x_16, x_2, x_17, x_20, x_72, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_73); -lean_dec(x_72); -return x_74; +lean_ctor_set(x_70, 0, x_69); +lean_ctor_set(x_70, 1, x_68); +x_71 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___closed__4; +x_72 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_72, 0, x_70); +lean_ctor_set(x_72, 1, x_71); +x_73 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_21, x_72, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_29); +x_74 = lean_ctor_get(x_73, 0); +lean_inc(x_74); +x_75 = lean_ctor_get(x_73, 1); +lean_inc(x_75); +lean_dec(x_73); +x_76 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___lambda__2(x_1, x_18, x_2, x_19, x_22, x_74, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_75); +lean_dec(x_74); +return x_76; } } } } else { -uint8_t x_87; +uint8_t x_89; +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_2); lean_dec(x_1); -x_87 = !lean_is_exclusive(x_12); -if (x_87 == 0) +x_89 = !lean_is_exclusive(x_14); +if (x_89 == 0) { -return x_12; +return x_14; } else { -lean_object* x_88; lean_object* x_89; lean_object* x_90; -x_88 = lean_ctor_get(x_12, 0); -x_89 = lean_ctor_get(x_12, 1); -lean_inc(x_89); -lean_inc(x_88); -lean_dec(x_12); -x_90 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_90, 0, x_88); -lean_ctor_set(x_90, 1, x_89); -return x_90; +lean_object* x_90; lean_object* x_91; lean_object* x_92; +x_90 = lean_ctor_get(x_14, 0); +x_91 = lean_ctor_get(x_14, 1); +lean_inc(x_91); +lean_inc(x_90); +lean_dec(x_14); +x_92 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_91); +return x_92; } } } @@ -15869,23 +16122,38 @@ lean_dec(x_2); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___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_Elab_Tactic_Omega_MetaProblem_addIntEquality___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_11; -x_11 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +uint8_t x_13; lean_object* x_14; +x_13 = lean_unbox(x_6); +lean_dec(x_6); +x_14 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___lambda__1(x_1, x_2, x_3, x_4, x_5, x_13, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_7); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -return x_11; +return x_14; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___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_Elab_Tactic_Omega_MetaProblem_addIntEquality___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_15; -x_15 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___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); +uint8_t x_17; lean_object* x_18; +x_17 = lean_unbox(x_10); +lean_dec(x_10); +x_18 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_17, x_11, x_12, x_13, x_14, x_15, x_16); lean_dec(x_6); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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_7); +lean_dec(x_7); +x_15 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality(x_1, x_2, x_3, x_4, x_5, x_6, x_14, x_8, x_9, x_10, x_11, x_12, x_13); return x_15; } } @@ -15968,189 +16236,189 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality___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_Elab_Tactic_Omega_MetaProblem_addIntInequality___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_object* x_11, lean_object* x_12) { _start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_11 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__3___closed__5; -x_12 = lean_array_push(x_11, x_1); -x_13 = lean_array_push(x_12, x_2); -x_14 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality___lambda__1___closed__2; -x_15 = l_Lean_Meta_mkAppM(x_14, x_13, x_6, x_7, x_8, x_9, x_10); -return x_15; +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_13 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__3___closed__5; +x_14 = lean_array_push(x_13, x_1); +x_15 = lean_array_push(x_14, x_2); +x_16 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality___lambda__1___closed__2; +x_17 = l_Lean_Meta_mkAppM(x_16, x_15, x_8, x_9, x_10, x_11, x_12); +return x_17; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality___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_Elab_Tactic_Omega_MetaProblem_addIntInequality___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, 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) { _start: { -uint8_t x_15; -x_15 = !lean_is_exclusive(x_1); -if (x_15 == 0) +uint8_t x_17; +x_17 = !lean_is_exclusive(x_1); +if (x_17 == 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; -x_16 = lean_ctor_get(x_1, 0); -x_17 = lean_ctor_get(x_1, 1); -x_18 = lean_ctor_get(x_1, 2); -x_19 = lean_ctor_get(x_1, 3); -x_20 = lean_ctor_get(x_2, 0); -lean_inc(x_20); -x_21 = lean_ctor_get(x_2, 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; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_18 = lean_ctor_get(x_1, 0); +x_19 = lean_ctor_get(x_1, 1); +x_20 = lean_ctor_get(x_1, 2); +x_21 = lean_ctor_get(x_1, 3); +x_22 = lean_ctor_get(x_2, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_2, 1); +lean_inc(x_23); lean_dec(x_2); -x_22 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality___lambda__1___boxed), 10, 1); -lean_closure_set(x_22, 0, x_3); -x_23 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg), 10, 2); -lean_closure_set(x_23, 0, x_4); -lean_closure_set(x_23, 1, x_22); -x_24 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_24, 0, x_23); -x_25 = l_Lean_Elab_Tactic_Omega_Problem_addInequality(x_16, x_20, x_21, x_24); -x_26 = l_Lean_Elab_Tactic_Omega_Problem_solveEqualities(x_25, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_26) == 0) +x_24 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality___lambda__1___boxed), 12, 1); +lean_closure_set(x_24, 0, x_3); +x_25 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg___boxed), 12, 2); +lean_closure_set(x_25, 0, x_4); +lean_closure_set(x_25, 1, x_24); +x_26 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_26, 0, x_25); +x_27 = l_Lean_Elab_Tactic_Omega_Problem_addInequality(x_18, x_22, x_23, x_26); +x_28 = l_Lean_Elab_Tactic_Omega_Problem_solveEqualities(x_27, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_28) == 0) { -uint8_t x_27; -x_27 = !lean_is_exclusive(x_26); -if (x_27 == 0) +uint8_t x_29; +x_29 = !lean_is_exclusive(x_28); +if (x_29 == 0) { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_26, 0); -x_29 = l_Lean_HashSet_toList___at_Lean_Elab_Tactic_Omega_lookup___spec__10(x_5); -x_30 = l_List_appendTR___rarg(x_29, x_17); -lean_ctor_set(x_1, 1, x_30); -lean_ctor_set(x_1, 0, x_28); -lean_ctor_set(x_26, 0, x_1); -return x_26; +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_28, 0); +x_31 = l_Lean_HashSet_toList___at_Lean_Elab_Tactic_Omega_lookup___spec__10(x_5); +x_32 = l_List_appendTR___rarg(x_31, x_19); +lean_ctor_set(x_1, 1, x_32); +lean_ctor_set(x_1, 0, x_30); +lean_ctor_set(x_28, 0, x_1); +return x_28; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_31 = lean_ctor_get(x_26, 0); -x_32 = lean_ctor_get(x_26, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_26); -x_33 = l_Lean_HashSet_toList___at_Lean_Elab_Tactic_Omega_lookup___spec__10(x_5); -x_34 = l_List_appendTR___rarg(x_33, x_17); -lean_ctor_set(x_1, 1, x_34); -lean_ctor_set(x_1, 0, x_31); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_1); -lean_ctor_set(x_35, 1, x_32); -return x_35; +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_33 = lean_ctor_get(x_28, 0); +x_34 = lean_ctor_get(x_28, 1); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_28); +x_35 = l_Lean_HashSet_toList___at_Lean_Elab_Tactic_Omega_lookup___spec__10(x_5); +x_36 = l_List_appendTR___rarg(x_35, x_19); +lean_ctor_set(x_1, 1, x_36); +lean_ctor_set(x_1, 0, x_33); +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_1); +lean_ctor_set(x_37, 1, x_34); +return x_37; } } else { -uint8_t x_36; +uint8_t x_38; lean_free_object(x_1); +lean_dec(x_21); +lean_dec(x_20); lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_17); lean_dec(x_5); -x_36 = !lean_is_exclusive(x_26); -if (x_36 == 0) +x_38 = !lean_is_exclusive(x_28); +if (x_38 == 0) { -return x_26; +return x_28; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_26, 0); -x_38 = lean_ctor_get(x_26, 1); -lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_26); -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_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_28, 0); +x_40 = lean_ctor_get(x_28, 1); +lean_inc(x_40); +lean_inc(x_39); +lean_dec(x_28); +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_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; 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_40 = lean_ctor_get(x_1, 0); -x_41 = lean_ctor_get(x_1, 1); -x_42 = lean_ctor_get(x_1, 2); -x_43 = lean_ctor_get(x_1, 3); +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; 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_42 = lean_ctor_get(x_1, 0); +x_43 = lean_ctor_get(x_1, 1); +x_44 = lean_ctor_get(x_1, 2); +x_45 = lean_ctor_get(x_1, 3); +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_1); -x_44 = lean_ctor_get(x_2, 0); -lean_inc(x_44); -x_45 = lean_ctor_get(x_2, 1); -lean_inc(x_45); +x_46 = lean_ctor_get(x_2, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_2, 1); +lean_inc(x_47); lean_dec(x_2); -x_46 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality___lambda__1___boxed), 10, 1); -lean_closure_set(x_46, 0, x_3); -x_47 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg), 10, 2); -lean_closure_set(x_47, 0, x_4); -lean_closure_set(x_47, 1, x_46); -x_48 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_48, 0, x_47); -x_49 = l_Lean_Elab_Tactic_Omega_Problem_addInequality(x_40, x_44, x_45, x_48); -x_50 = l_Lean_Elab_Tactic_Omega_Problem_solveEqualities(x_49, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -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; -x_51 = lean_ctor_get(x_50, 0); -lean_inc(x_51); -x_52 = lean_ctor_get(x_50, 1); -lean_inc(x_52); -if (lean_is_exclusive(x_50)) { - lean_ctor_release(x_50, 0); - lean_ctor_release(x_50, 1); - x_53 = x_50; +x_48 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality___lambda__1___boxed), 12, 1); +lean_closure_set(x_48, 0, x_3); +x_49 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg___boxed), 12, 2); +lean_closure_set(x_49, 0, x_4); +lean_closure_set(x_49, 1, x_48); +x_50 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_50, 0, x_49); +x_51 = l_Lean_Elab_Tactic_Omega_Problem_addInequality(x_42, x_46, x_47, x_50); +x_52 = l_Lean_Elab_Tactic_Omega_Problem_solveEqualities(x_51, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_52) == 0) +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_53 = lean_ctor_get(x_52, 0); +lean_inc(x_53); +x_54 = lean_ctor_get(x_52, 1); +lean_inc(x_54); +if (lean_is_exclusive(x_52)) { + lean_ctor_release(x_52, 0); + lean_ctor_release(x_52, 1); + x_55 = x_52; } else { - lean_dec_ref(x_50); - x_53 = lean_box(0); + lean_dec_ref(x_52); + x_55 = lean_box(0); } -x_54 = l_Lean_HashSet_toList___at_Lean_Elab_Tactic_Omega_lookup___spec__10(x_5); -x_55 = l_List_appendTR___rarg(x_54, x_41); -x_56 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_56, 0, x_51); -lean_ctor_set(x_56, 1, x_55); -lean_ctor_set(x_56, 2, x_42); -lean_ctor_set(x_56, 3, x_43); -if (lean_is_scalar(x_53)) { - x_57 = lean_alloc_ctor(0, 2, 0); +x_56 = l_Lean_HashSet_toList___at_Lean_Elab_Tactic_Omega_lookup___spec__10(x_5); +x_57 = l_List_appendTR___rarg(x_56, x_43); +x_58 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_58, 0, x_53); +lean_ctor_set(x_58, 1, x_57); +lean_ctor_set(x_58, 2, x_44); +lean_ctor_set(x_58, 3, x_45); +if (lean_is_scalar(x_55)) { + x_59 = lean_alloc_ctor(0, 2, 0); } else { - x_57 = x_53; + x_59 = x_55; } -lean_ctor_set(x_57, 0, x_56); -lean_ctor_set(x_57, 1, x_52); -return x_57; +lean_ctor_set(x_59, 0, x_58); +lean_ctor_set(x_59, 1, x_54); +return x_59; } else { -lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; +lean_dec(x_45); +lean_dec(x_44); lean_dec(x_43); -lean_dec(x_42); -lean_dec(x_41); lean_dec(x_5); -x_58 = lean_ctor_get(x_50, 0); -lean_inc(x_58); -x_59 = lean_ctor_get(x_50, 1); -lean_inc(x_59); -if (lean_is_exclusive(x_50)) { - lean_ctor_release(x_50, 0); - lean_ctor_release(x_50, 1); - x_60 = x_50; +x_60 = lean_ctor_get(x_52, 0); +lean_inc(x_60); +x_61 = lean_ctor_get(x_52, 1); +lean_inc(x_61); +if (lean_is_exclusive(x_52)) { + lean_ctor_release(x_52, 0); + lean_ctor_release(x_52, 1); + x_62 = x_52; } else { - lean_dec_ref(x_50); - x_60 = lean_box(0); + lean_dec_ref(x_52); + x_62 = lean_box(0); } -if (lean_is_scalar(x_60)) { - x_61 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_62)) { + x_63 = lean_alloc_ctor(1, 2, 0); } else { - x_61 = x_60; + x_63 = x_62; } -lean_ctor_set(x_61, 0, x_58); -lean_ctor_set(x_61, 1, x_59); -return x_61; +lean_ctor_set(x_63, 0, x_60); +lean_ctor_set(x_63, 1, x_61); +return x_63; } } } @@ -16172,224 +16440,226 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_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_Omega_MetaProblem_addIntInequality(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_object* x_12, lean_object* x_13) { _start: { -lean_object* x_12; +lean_object* 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_6); lean_inc(x_5); lean_inc(x_4); -x_12 = l_Lean_Elab_Tactic_Omega_asLinearCombo(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_12) == 0) +x_14 = l_Lean_Elab_Tactic_Omega_asLinearCombo(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_14) == 0) { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_76; lean_object* x_77; lean_object* x_78; uint8_t x_79; -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_13, 1); -lean_inc(x_14); -x_15 = lean_ctor_get(x_12, 1); +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_78; lean_object* x_79; lean_object* x_80; uint8_t x_81; +x_15 = lean_ctor_get(x_14, 0); lean_inc(x_15); -lean_dec(x_12); -x_16 = lean_ctor_get(x_13, 0); +x_16 = lean_ctor_get(x_15, 1); lean_inc(x_16); -lean_dec(x_13); -x_17 = lean_ctor_get(x_14, 0); +x_17 = lean_ctor_get(x_14, 1); lean_inc(x_17); -x_18 = lean_ctor_get(x_14, 1); -lean_inc(x_18); lean_dec(x_14); -x_19 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__2; -x_76 = lean_ctor_get(x_18, 1); -lean_inc(x_76); -lean_dec(x_18); -x_77 = lean_array_get_size(x_76); -x_78 = lean_unsigned_to_nat(0u); -x_79 = lean_nat_dec_lt(x_78, x_77); -if (x_79 == 0) +x_18 = lean_ctor_get(x_15, 0); +lean_inc(x_18); +lean_dec(x_15); +x_19 = lean_ctor_get(x_16, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_16, 1); +lean_inc(x_20); +lean_dec(x_16); +x_21 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__2; +x_78 = lean_ctor_get(x_20, 1); +lean_inc(x_78); +lean_dec(x_20); +x_79 = lean_array_get_size(x_78); +x_80 = lean_unsigned_to_nat(0u); +x_81 = lean_nat_dec_lt(x_80, x_79); +if (x_81 == 0) { -lean_object* x_80; -lean_dec(x_77); -lean_dec(x_76); -x_80 = l_Lean_Elab_Tactic_Omega_MetaProblem_problem___default___closed__2; -x_20 = x_80; -goto block_75; +lean_object* x_82; +lean_dec(x_79); +lean_dec(x_78); +x_82 = l_Lean_Elab_Tactic_Omega_MetaProblem_problem___default___closed__2; +x_22 = x_82; +goto block_77; } else { -uint8_t x_81; -x_81 = lean_nat_dec_le(x_77, x_77); -if (x_81 == 0) +uint8_t x_83; +x_83 = lean_nat_dec_le(x_79, x_79); +if (x_83 == 0) { -lean_object* x_82; -lean_dec(x_77); -lean_dec(x_76); -x_82 = l_Lean_Elab_Tactic_Omega_MetaProblem_problem___default___closed__2; -x_20 = x_82; -goto block_75; +lean_object* x_84; +lean_dec(x_79); +lean_dec(x_78); +x_84 = l_Lean_Elab_Tactic_Omega_MetaProblem_problem___default___closed__2; +x_22 = x_84; +goto block_77; } else { -size_t x_83; size_t x_84; lean_object* x_85; lean_object* x_86; -x_83 = 0; -x_84 = lean_usize_of_nat(x_77); -lean_dec(x_77); -x_85 = l_Lean_Elab_Tactic_Omega_MetaProblem_problem___default___closed__2; +size_t x_85; size_t x_86; lean_object* x_87; lean_object* x_88; +x_85 = 0; +x_86 = lean_usize_of_nat(x_79); +lean_dec(x_79); +x_87 = l_Lean_Elab_Tactic_Omega_MetaProblem_problem___default___closed__2; lean_inc(x_1); -x_86 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality___spec__2(x_1, x_76, x_83, x_84, x_85); -lean_dec(x_76); -x_20 = x_86; -goto block_75; +x_88 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality___spec__2(x_1, x_78, x_85, x_86, x_87); +lean_dec(x_78); +x_22 = x_88; +goto block_77; } } -block_75: -{ -lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_21 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_19, x_4, x_5, x_6, x_7, x_8, x_9, x_10, 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) +block_77: { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_21, 1); +lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_23 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_21, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_17); +x_24 = lean_ctor_get(x_23, 0); lean_inc(x_24); -lean_dec(x_21); -x_25 = lean_box(0); -x_26 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality___lambda__2(x_1, x_16, x_2, x_17, x_20, x_25, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_24); -return x_26; +x_25 = lean_unbox(x_24); +lean_dec(x_24); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_23, 1); +lean_inc(x_26); +lean_dec(x_23); +x_27 = lean_box(0); +x_28 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality___lambda__2(x_1, x_18, x_2, x_19, x_22, x_27, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_26); +return x_28; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_27 = lean_ctor_get(x_21, 1); -lean_inc(x_27); -lean_dec(x_21); -x_28 = lean_ctor_get(x_16, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_16, 1); +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; 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_29 = lean_ctor_get(x_23, 1); lean_inc(x_29); -x_30 = lean_unsigned_to_nat(0u); -x_31 = l_List_enumFrom___rarg(x_30, x_29); -x_32 = lean_box(0); -x_33 = l_List_mapTR_loop___at_Lean_Omega_LinearCombo_instToStringLinearCombo___spec__1(x_31, x_32); -x_34 = l_Lean_Elab_Tactic_Omega_MetaProblem_problem___default___lambda__1___closed__1; -x_35 = l_List_foldl___at_String_join___spec__1(x_34, x_33); -lean_dec(x_33); -x_36 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__8; -x_37 = lean_int_dec_lt(x_28, x_36); -if (x_37 == 0) -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; -x_38 = lean_nat_abs(x_28); -lean_dec(x_28); -x_39 = l_Nat_repr(x_38); -x_40 = lean_string_append(x_34, x_39); -lean_dec(x_39); -x_41 = lean_string_append(x_40, x_34); -x_42 = lean_string_append(x_41, x_35); +lean_dec(x_23); +x_30 = lean_ctor_get(x_18, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_18, 1); +lean_inc(x_31); +x_32 = lean_unsigned_to_nat(0u); +x_33 = l_List_enumFrom___rarg(x_32, x_31); +x_34 = lean_box(0); +x_35 = l_List_mapTR_loop___at_Lean_Omega_LinearCombo_instToStringLinearCombo___spec__1(x_33, x_34); +x_36 = l_Lean_Elab_Tactic_Omega_MetaProblem_problem___default___lambda__1___closed__1; +x_37 = l_List_foldl___at_String_join___spec__1(x_36, x_35); lean_dec(x_35); -x_43 = lean_string_append(x_42, x_34); -x_44 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_44, 0, x_43); -x_45 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_45, 0, x_44); -x_46 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___closed__2; -x_47 = lean_alloc_ctor(7, 2, 0); +x_38 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__8; +x_39 = lean_int_dec_lt(x_30, x_38); +if (x_39 == 0) +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_40 = lean_nat_abs(x_30); +lean_dec(x_30); +x_41 = l_Nat_repr(x_40); +x_42 = lean_string_append(x_36, x_41); +lean_dec(x_41); +x_43 = lean_string_append(x_42, x_36); +x_44 = lean_string_append(x_43, x_37); +lean_dec(x_37); +x_45 = lean_string_append(x_44, x_36); +x_46 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_46, 0, x_45); +x_47 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_45); -x_48 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality___closed__2; +x_48 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___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 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_19, x_49, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_27); -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 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality___lambda__2(x_1, x_16, x_2, x_17, x_20, x_51, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_52); -lean_dec(x_51); -return x_53; +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_47); +x_50 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality___closed__2; +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 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_21, x_51, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_29); +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_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality___lambda__2(x_1, x_18, x_2, x_19, x_22, x_53, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_54); +lean_dec(x_53); +return x_55; } else { -lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; -x_54 = lean_nat_abs(x_28); -lean_dec(x_28); -x_55 = lean_unsigned_to_nat(1u); -x_56 = lean_nat_sub(x_54, x_55); -lean_dec(x_54); -x_57 = lean_nat_add(x_56, x_55); +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_56 = lean_nat_abs(x_30); +lean_dec(x_30); +x_57 = lean_unsigned_to_nat(1u); +x_58 = lean_nat_sub(x_56, x_57); lean_dec(x_56); -x_58 = l_Nat_repr(x_57); -x_59 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___closed__5; -x_60 = lean_string_append(x_59, x_58); +x_59 = lean_nat_add(x_58, x_57); lean_dec(x_58); -x_61 = lean_string_append(x_34, x_60); +x_60 = l_Nat_repr(x_59); +x_61 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___closed__5; +x_62 = lean_string_append(x_61, x_60); lean_dec(x_60); -x_62 = lean_string_append(x_61, x_34); -x_63 = lean_string_append(x_62, x_35); -lean_dec(x_35); -x_64 = lean_string_append(x_63, x_34); -x_65 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_65, 0, x_64); -x_66 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_66, 0, x_65); -x_67 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___closed__2; -x_68 = lean_alloc_ctor(7, 2, 0); +x_63 = lean_string_append(x_36, x_62); +lean_dec(x_62); +x_64 = lean_string_append(x_63, x_36); +x_65 = lean_string_append(x_64, x_37); +lean_dec(x_37); +x_66 = lean_string_append(x_65, x_36); +x_67 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_67, 0, x_66); +x_68 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_68, 0, x_67); -lean_ctor_set(x_68, 1, x_66); -x_69 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality___closed__2; +x_69 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___closed__2; x_70 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_70, 0, x_68); -lean_ctor_set(x_70, 1, x_69); -x_71 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_19, x_70, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_27); -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 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality___lambda__2(x_1, x_16, x_2, x_17, x_20, x_72, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_73); -lean_dec(x_72); -return x_74; +lean_ctor_set(x_70, 0, x_69); +lean_ctor_set(x_70, 1, x_68); +x_71 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality___closed__2; +x_72 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_72, 0, x_70); +lean_ctor_set(x_72, 1, x_71); +x_73 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_21, x_72, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_29); +x_74 = lean_ctor_get(x_73, 0); +lean_inc(x_74); +x_75 = lean_ctor_get(x_73, 1); +lean_inc(x_75); +lean_dec(x_73); +x_76 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality___lambda__2(x_1, x_18, x_2, x_19, x_22, x_74, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_75); +lean_dec(x_74); +return x_76; } } } } else { -uint8_t x_87; +uint8_t x_89; +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_2); lean_dec(x_1); -x_87 = !lean_is_exclusive(x_12); -if (x_87 == 0) +x_89 = !lean_is_exclusive(x_14); +if (x_89 == 0) { -return x_12; +return x_14; } else { -lean_object* x_88; lean_object* x_89; lean_object* x_90; -x_88 = lean_ctor_get(x_12, 0); -x_89 = lean_ctor_get(x_12, 1); -lean_inc(x_89); -lean_inc(x_88); -lean_dec(x_12); -x_90 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_90, 0, x_88); -lean_ctor_set(x_90, 1, x_89); -return x_90; +lean_object* x_90; lean_object* x_91; lean_object* x_92; +x_90 = lean_ctor_get(x_14, 0); +x_91 = lean_ctor_get(x_14, 1); +lean_inc(x_91); +lean_inc(x_90); +lean_dec(x_14); +x_92 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_91); +return x_92; } } } @@ -16407,23 +16677,38 @@ lean_dec(x_2); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality___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_Elab_Tactic_Omega_MetaProblem_addIntInequality___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_11; -x_11 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +uint8_t x_13; lean_object* x_14; +x_13 = lean_unbox(x_6); +lean_dec(x_6); +x_14 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality___lambda__1(x_1, x_2, x_3, x_4, x_5, x_13, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_7); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -return x_11; +return x_14; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality___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_Elab_Tactic_Omega_MetaProblem_addIntInequality___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_15; -x_15 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality___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); +uint8_t x_17; lean_object* x_18; +x_17 = lean_unbox(x_10); +lean_dec(x_10); +x_18 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_17, x_11, x_12, x_13, x_14, x_15, x_16); lean_dec(x_6); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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_7); +lean_dec(x_7); +x_15 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality(x_1, x_2, x_3, x_4, x_5, x_6, x_14, x_8, x_9, x_10, x_11, x_12, x_13); return x_15; } } @@ -18487,106 +18772,106 @@ lean_dec(x_4); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at_Lean_Elab_Tactic_Omega_MetaProblem_addFact___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_EXPORT lean_object* l_Lean_instantiateMVars___at_Lean_Elab_Tactic_Omega_MetaProblem_addFact___spec__1(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) { _start: { -uint8_t x_10; -x_10 = l_Lean_Expr_hasMVar(x_1); -if (x_10 == 0) +uint8_t x_12; +x_12 = l_Lean_Expr_hasMVar(x_1); +if (x_12 == 0) { -lean_object* x_11; -x_11 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_11, 0, x_1); -lean_ctor_set(x_11, 1, x_9); -return x_11; +lean_object* x_13; +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_1); +lean_ctor_set(x_13, 1, x_11); +return x_13; } else { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_12 = lean_st_ref_get(x_6, x_9); -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_12, 1); -lean_inc(x_14); -lean_dec(x_12); -x_15 = lean_ctor_get(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; uint8_t x_24; +x_14 = lean_st_ref_get(x_8, x_11); +x_15 = lean_ctor_get(x_14, 0); lean_inc(x_15); -lean_dec(x_13); -x_16 = l_Lean_instantiateMVarsCore(x_15, x_1); -x_17 = lean_ctor_get(x_16, 0); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = lean_ctor_get(x_15, 0); lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); -lean_inc(x_18); -lean_dec(x_16); -x_19 = lean_st_ref_take(x_6, x_14); -x_20 = lean_ctor_get(x_19, 0); +lean_dec(x_15); +x_18 = l_Lean_instantiateMVarsCore(x_17, 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); -x_21 = lean_ctor_get(x_19, 1); -lean_inc(x_21); -lean_dec(x_19); -x_22 = !lean_is_exclusive(x_20); -if (x_22 == 0) +lean_dec(x_18); +x_21 = lean_st_ref_take(x_8, x_16); +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_23; lean_object* x_24; uint8_t x_25; -x_23 = lean_ctor_get(x_20, 0); -lean_dec(x_23); -lean_ctor_set(x_20, 0, x_18); -x_24 = lean_st_ref_set(x_6, x_20, x_21); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_25 = lean_ctor_get(x_22, 0); +lean_dec(x_25); +lean_ctor_set(x_22, 0, x_20); +x_26 = lean_st_ref_set(x_8, x_22, x_23); +x_27 = !lean_is_exclusive(x_26); +if (x_27 == 0) { -lean_object* x_26; -x_26 = lean_ctor_get(x_24, 0); -lean_dec(x_26); -lean_ctor_set(x_24, 0, x_17); -return x_24; +lean_object* x_28; +x_28 = lean_ctor_get(x_26, 0); +lean_dec(x_28); +lean_ctor_set(x_26, 0, x_19); +return x_26; } else { -lean_object* x_27; lean_object* x_28; -x_27 = lean_ctor_get(x_24, 1); -lean_inc(x_27); -lean_dec(x_24); -x_28 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_28, 0, x_17); -lean_ctor_set(x_28, 1, x_27); -return x_28; +lean_object* x_29; lean_object* x_30; +x_29 = lean_ctor_get(x_26, 1); +lean_inc(x_29); +lean_dec(x_26); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_19); +lean_ctor_set(x_30, 1, x_29); +return x_30; } } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_29 = lean_ctor_get(x_20, 1); -x_30 = lean_ctor_get(x_20, 2); -x_31 = lean_ctor_get(x_20, 3); +lean_object* x_31; lean_object* x_32; 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_31 = lean_ctor_get(x_22, 1); +x_32 = lean_ctor_get(x_22, 2); +x_33 = lean_ctor_get(x_22, 3); +lean_inc(x_33); +lean_inc(x_32); lean_inc(x_31); -lean_inc(x_30); -lean_inc(x_29); -lean_dec(x_20); -x_32 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_32, 0, x_18); -lean_ctor_set(x_32, 1, x_29); -lean_ctor_set(x_32, 2, x_30); -lean_ctor_set(x_32, 3, x_31); -x_33 = lean_st_ref_set(x_6, x_32, x_21); -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; +lean_dec(x_22); +x_34 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_34, 0, x_20); +lean_ctor_set(x_34, 1, x_31); +lean_ctor_set(x_34, 2, x_32); +lean_ctor_set(x_34, 3, x_33); +x_35 = lean_st_ref_set(x_8, x_34, x_23); +x_36 = lean_ctor_get(x_35, 1); +lean_inc(x_36); +if (lean_is_exclusive(x_35)) { + lean_ctor_release(x_35, 0); + lean_ctor_release(x_35, 1); + x_37 = x_35; } else { - lean_dec_ref(x_33); - x_35 = lean_box(0); + lean_dec_ref(x_35); + x_37 = lean_box(0); } -if (lean_is_scalar(x_35)) { - x_36 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_37)) { + x_38 = lean_alloc_ctor(0, 2, 0); } else { - x_36 = x_35; + x_38 = x_37; } -lean_ctor_set(x_36, 0, x_17); -lean_ctor_set(x_36, 1, x_34); -return x_36; +lean_ctor_set(x_38, 0, x_19); +lean_ctor_set(x_38, 1, x_36); +return x_38; } } } @@ -19291,184 +19576,185 @@ x_3 = l_Lean_Expr_const___override(x_2, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, uint8_t x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { lean_dec(x_5); switch (lean_obj_tag(x_1)) { case 5: { -lean_object* x_14; lean_object* x_15; -x_14 = l_Lean_Expr_getAppFnArgs(x_1); -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -if (lean_obj_tag(x_15) == 1) +lean_object* x_16; lean_object* x_17; +x_16 = l_Lean_Expr_getAppFnArgs(x_1); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +if (lean_obj_tag(x_17) == 1) { -lean_object* x_16; -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -switch (lean_obj_tag(x_16)) { +lean_object* x_18; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +switch (lean_obj_tag(x_18)) { case 0: { -uint8_t x_17; -x_17 = !lean_is_exclusive(x_14); -if (x_17 == 0) -{ -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_14, 1); -x_19 = lean_ctor_get(x_14, 0); -lean_dec(x_19); -x_20 = lean_ctor_get(x_15, 1); -lean_inc(x_20); -lean_dec(x_15); -x_21 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___lambda__2___closed__1; -x_22 = lean_string_dec_eq(x_20, x_21); -if (x_22 == 0) +uint8_t x_19; +x_19 = !lean_is_exclusive(x_16); +if (x_19 == 0) { -lean_object* x_23; uint8_t x_24; -x_23 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__1; -x_24 = lean_string_dec_eq(x_20, x_23); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = lean_ctor_get(x_16, 1); +x_21 = lean_ctor_get(x_16, 0); +lean_dec(x_21); +x_22 = lean_ctor_get(x_17, 1); +lean_inc(x_22); +lean_dec(x_17); +x_23 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___lambda__2___closed__1; +x_24 = lean_string_dec_eq(x_22, x_23); if (x_24 == 0) { lean_object* x_25; uint8_t x_26; -x_25 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__1; -x_26 = lean_string_dec_eq(x_20, x_25); +x_25 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__1; +x_26 = lean_string_dec_eq(x_22, x_25); if (x_26 == 0) { lean_object* x_27; uint8_t x_28; -x_27 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__5; -x_28 = lean_string_dec_eq(x_20, x_27); +x_27 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__1; +x_28 = lean_string_dec_eq(x_22, x_27); if (x_28 == 0) { lean_object* x_29; uint8_t x_30; -x_29 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__2; -x_30 = lean_string_dec_eq(x_20, x_29); +x_29 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__5; +x_30 = lean_string_dec_eq(x_22, x_29); if (x_30 == 0) { lean_object* x_31; uint8_t x_32; -x_31 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__3; -x_32 = lean_string_dec_eq(x_20, x_31); +x_31 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__2; +x_32 = lean_string_dec_eq(x_22, x_31); if (x_32 == 0) { lean_object* x_33; uint8_t x_34; -x_33 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__3; -x_34 = lean_string_dec_eq(x_20, x_33); +x_33 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__3; +x_34 = lean_string_dec_eq(x_22, x_33); if (x_34 == 0) { lean_object* x_35; uint8_t x_36; +x_35 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__3; +x_36 = lean_string_dec_eq(x_22, x_35); +if (x_36 == 0) +{ +lean_object* x_37; uint8_t x_38; +lean_dec(x_14); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); -x_35 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__4___closed__1; -x_36 = lean_string_dec_eq(x_20, x_35); -lean_dec(x_20); -if (x_36 == 0) +x_37 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__4___closed__1; +x_38 = lean_string_dec_eq(x_22, x_37); +lean_dec(x_22); +if (x_38 == 0) { -lean_object* x_37; lean_object* x_38; -lean_dec(x_18); +lean_object* x_39; lean_object* x_40; +lean_dec(x_20); lean_dec(x_8); lean_dec(x_4); lean_dec(x_3); -x_37 = lean_unsigned_to_nat(0u); -lean_ctor_set(x_14, 1, x_37); -lean_ctor_set(x_14, 0, x_2); -x_38 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_38, 0, x_14); -lean_ctor_set(x_38, 1, x_13); -return x_38; +x_39 = lean_unsigned_to_nat(0u); +lean_ctor_set(x_16, 1, x_39); +lean_ctor_set(x_16, 0, x_2); +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_16); +lean_ctor_set(x_40, 1, x_15); +return x_40; } else { -lean_object* x_39; lean_object* x_40; uint8_t x_41; -x_39 = lean_array_get_size(x_18); -lean_dec(x_18); -x_40 = lean_unsigned_to_nat(2u); -x_41 = lean_nat_dec_eq(x_39, x_40); -lean_dec(x_39); -if (x_41 == 0) +lean_object* x_41; lean_object* x_42; uint8_t x_43; +x_41 = lean_array_get_size(x_20); +lean_dec(x_20); +x_42 = lean_unsigned_to_nat(2u); +x_43 = lean_nat_dec_eq(x_41, x_42); +lean_dec(x_41); +if (x_43 == 0) { -lean_object* x_42; lean_object* x_43; +lean_object* x_44; lean_object* x_45; lean_dec(x_8); lean_dec(x_4); lean_dec(x_3); -x_42 = lean_unsigned_to_nat(0u); -lean_ctor_set(x_14, 1, x_42); -lean_ctor_set(x_14, 0, x_2); -x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_14); -lean_ctor_set(x_43, 1, x_13); -return x_43; +x_44 = lean_unsigned_to_nat(0u); +lean_ctor_set(x_16, 1, x_44); +lean_ctor_set(x_16, 0, x_2); +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_16); +lean_ctor_set(x_45, 1, x_15); +return x_45; } else { -uint8_t x_44; -x_44 = lean_ctor_get_uint8(x_8, 0); +uint8_t x_46; +x_46 = lean_ctor_get_uint8(x_8, 0); lean_dec(x_8); -if (x_44 == 0) +if (x_46 == 0) { -lean_object* x_45; lean_object* x_46; +lean_object* x_47; lean_object* x_48; lean_dec(x_4); lean_dec(x_3); -x_45 = lean_unsigned_to_nat(0u); -lean_ctor_set(x_14, 1, x_45); -lean_ctor_set(x_14, 0, x_2); -x_46 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_46, 0, x_14); -lean_ctor_set(x_46, 1, x_13); -return x_46; +x_47 = lean_unsigned_to_nat(0u); +lean_ctor_set(x_16, 1, x_47); +lean_ctor_set(x_16, 0, x_2); +x_48 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_48, 0, x_16); +lean_ctor_set(x_48, 1, x_15); +return x_48; } else { -lean_object* x_47; lean_object* x_48; uint8_t x_49; -x_47 = lean_ctor_get(x_2, 1); -lean_inc(x_47); -x_48 = lean_ctor_get(x_2, 2); -lean_inc(x_48); -x_49 = l_List_elem___at_Lean_CollectLevelParams_visitExpr___spec__2(x_3, x_48); -if (x_49 == 0) -{ -lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_50 = lean_ctor_get(x_2, 3); +lean_object* x_49; lean_object* x_50; uint8_t x_51; +x_49 = lean_ctor_get(x_2, 1); +lean_inc(x_49); +x_50 = lean_ctor_get(x_2, 2); lean_inc(x_50); +x_51 = l_List_elem___at_Lean_CollectLevelParams_visitExpr___spec__2(x_3, x_50); +if (x_51 == 0) +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_52 = lean_ctor_get(x_2, 3); +lean_inc(x_52); lean_dec(x_2); -x_51 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_51, 0, x_3); -lean_ctor_set(x_51, 1, x_48); -x_52 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_52, 0, x_4); -lean_ctor_set(x_52, 1, x_47); -lean_ctor_set(x_52, 2, x_51); -lean_ctor_set(x_52, 3, x_50); -x_53 = lean_unsigned_to_nat(1u); -lean_ctor_set(x_14, 1, x_53); -lean_ctor_set(x_14, 0, x_52); -x_54 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_54, 0, x_14); -lean_ctor_set(x_54, 1, x_13); -return x_54; +x_53 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_53, 0, x_3); +lean_ctor_set(x_53, 1, x_50); +x_54 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_54, 0, x_4); +lean_ctor_set(x_54, 1, x_49); +lean_ctor_set(x_54, 2, x_53); +lean_ctor_set(x_54, 3, x_52); +x_55 = lean_unsigned_to_nat(1u); +lean_ctor_set(x_16, 1, x_55); +lean_ctor_set(x_16, 0, x_54); +x_56 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_56, 0, x_16); +lean_ctor_set(x_56, 1, x_15); +return x_56; } else { -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_dec(x_3); -x_55 = lean_ctor_get(x_2, 3); -lean_inc(x_55); +x_57 = lean_ctor_get(x_2, 3); +lean_inc(x_57); lean_dec(x_2); -x_56 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_56, 0, x_4); -lean_ctor_set(x_56, 1, x_47); -lean_ctor_set(x_56, 2, x_48); -lean_ctor_set(x_56, 3, x_55); -x_57 = lean_unsigned_to_nat(1u); -lean_ctor_set(x_14, 1, x_57); -lean_ctor_set(x_14, 0, x_56); -x_58 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_58, 0, x_14); -lean_ctor_set(x_58, 1, x_13); -return x_58; +x_58 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_58, 0, x_4); +lean_ctor_set(x_58, 1, x_49); +lean_ctor_set(x_58, 2, x_50); +lean_ctor_set(x_58, 3, x_57); +x_59 = lean_unsigned_to_nat(1u); +lean_ctor_set(x_16, 1, x_59); +lean_ctor_set(x_16, 0, x_58); +x_60 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_60, 0, x_16); +lean_ctor_set(x_60, 1, x_15); +return x_60; } } } @@ -19476,145 +19762,148 @@ return x_58; } else { -lean_object* x_59; lean_object* x_60; uint8_t x_61; -lean_dec(x_20); +lean_object* x_61; lean_object* x_62; uint8_t x_63; +lean_dec(x_22); lean_dec(x_4); -x_59 = lean_array_get_size(x_18); -x_60 = lean_unsigned_to_nat(2u); -x_61 = lean_nat_dec_eq(x_59, x_60); -lean_dec(x_59); -if (x_61 == 0) +x_61 = lean_array_get_size(x_20); +x_62 = lean_unsigned_to_nat(2u); +x_63 = lean_nat_dec_eq(x_61, x_62); +lean_dec(x_61); +if (x_63 == 0) { -lean_object* x_62; lean_object* x_63; -lean_dec(x_18); +lean_object* x_64; lean_object* x_65; +lean_dec(x_20); +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_3); -x_62 = lean_unsigned_to_nat(0u); -lean_ctor_set(x_14, 1, x_62); -lean_ctor_set(x_14, 0, x_2); -x_63 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_63, 0, x_14); -lean_ctor_set(x_63, 1, x_13); -return x_63; +x_64 = lean_unsigned_to_nat(0u); +lean_ctor_set(x_16, 1, x_64); +lean_ctor_set(x_16, 0, x_2); +x_65 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_65, 0, x_16); +lean_ctor_set(x_65, 1, x_15); +return x_65; } else { -lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; -lean_free_object(x_14); -x_64 = lean_unsigned_to_nat(0u); -x_65 = lean_array_fget(x_18, x_64); -x_66 = lean_unsigned_to_nat(1u); -x_67 = lean_array_fget(x_18, x_66); -lean_dec(x_18); -x_68 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__1___closed__8; -lean_inc(x_67); -x_69 = l_Lean_Expr_app___override(x_68, x_67); -x_70 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__6; -x_71 = l_Lean_mkApp4(x_70, x_65, x_67, x_69, x_3); -x_72 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_71, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_72; +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_free_object(x_16); +x_66 = lean_unsigned_to_nat(0u); +x_67 = lean_array_fget(x_20, x_66); +x_68 = lean_unsigned_to_nat(1u); +x_69 = lean_array_fget(x_20, x_68); +lean_dec(x_20); +x_70 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__1___closed__8; +lean_inc(x_69); +x_71 = l_Lean_Expr_app___override(x_70, x_69); +x_72 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__6; +x_73 = l_Lean_mkApp4(x_72, x_67, x_69, x_71, x_3); +x_74 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_73, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_74; } } } else { -lean_object* x_73; lean_object* x_74; uint8_t x_75; -lean_dec(x_20); +lean_object* x_75; lean_object* x_76; uint8_t x_77; +lean_dec(x_22); lean_dec(x_4); -x_73 = lean_array_get_size(x_18); -x_74 = lean_unsigned_to_nat(2u); -x_75 = lean_nat_dec_eq(x_73, x_74); -lean_dec(x_73); -if (x_75 == 0) +x_75 = lean_array_get_size(x_20); +x_76 = lean_unsigned_to_nat(2u); +x_77 = lean_nat_dec_eq(x_75, x_76); +lean_dec(x_75); +if (x_77 == 0) { -lean_object* x_76; lean_object* x_77; -lean_dec(x_18); +lean_object* x_78; lean_object* x_79; +lean_dec(x_20); +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_3); -x_76 = lean_unsigned_to_nat(0u); -lean_ctor_set(x_14, 1, x_76); -lean_ctor_set(x_14, 0, x_2); -x_77 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_77, 0, x_14); -lean_ctor_set(x_77, 1, x_13); -return x_77; +x_78 = lean_unsigned_to_nat(0u); +lean_ctor_set(x_16, 1, x_78); +lean_ctor_set(x_16, 0, x_2); +x_79 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_79, 0, x_16); +lean_ctor_set(x_79, 1, x_15); +return x_79; } else { -lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; -lean_free_object(x_14); -x_78 = lean_unsigned_to_nat(0u); -x_79 = lean_array_fget(x_18, x_78); -x_80 = lean_unsigned_to_nat(1u); -x_81 = lean_array_fget(x_18, x_80); -lean_dec(x_18); +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; +lean_free_object(x_16); +x_80 = lean_unsigned_to_nat(0u); +x_81 = lean_array_fget(x_20, x_80); +x_82 = lean_unsigned_to_nat(1u); +x_83 = lean_array_fget(x_20, x_82); +lean_dec(x_20); +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_79); -x_82 = l_Lean_Meta_getLevel(x_79, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_82) == 0) +lean_inc(x_81); +x_84 = l_Lean_Meta_getLevel(x_81, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_84) == 0) { -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_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_box(0); -x_86 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_86, 0, x_83); -lean_ctor_set(x_86, 1, x_85); -x_87 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__8; -x_88 = l_Lean_Expr_const___override(x_87, x_86); -x_89 = l_Lean_mkApp3(x_88, x_79, x_81, x_3); -x_90 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_89, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_84); -return x_90; +lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; +x_85 = lean_ctor_get(x_84, 0); +lean_inc(x_85); +x_86 = lean_ctor_get(x_84, 1); +lean_inc(x_86); +lean_dec(x_84); +x_87 = lean_box(0); +x_88 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_88, 0, x_85); +lean_ctor_set(x_88, 1, x_87); +x_89 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__8; +x_90 = l_Lean_Expr_const___override(x_89, x_88); +x_91 = l_Lean_mkApp3(x_90, x_81, x_83, x_3); +x_92 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_91, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_86); +return x_92; } else { -uint8_t x_91; +uint8_t x_93; +lean_dec(x_83); lean_dec(x_81); -lean_dec(x_79); +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_3); lean_dec(x_2); -x_91 = !lean_is_exclusive(x_82); -if (x_91 == 0) +x_93 = !lean_is_exclusive(x_84); +if (x_93 == 0) { -return x_82; +return x_84; } else { -lean_object* x_92; lean_object* x_93; lean_object* x_94; -x_92 = lean_ctor_get(x_82, 0); -x_93 = lean_ctor_get(x_82, 1); -lean_inc(x_93); -lean_inc(x_92); -lean_dec(x_82); -x_94 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_94, 0, x_92); -lean_ctor_set(x_94, 1, x_93); -return x_94; +lean_object* x_94; lean_object* x_95; lean_object* x_96; +x_94 = lean_ctor_get(x_84, 0); +x_95 = lean_ctor_get(x_84, 1); +lean_inc(x_95); +lean_inc(x_94); +lean_dec(x_84); +x_96 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_96, 0, x_94); +lean_ctor_set(x_96, 1, x_95); +return x_96; } } } @@ -19622,97 +19911,99 @@ return x_94; } else { -lean_object* x_95; lean_object* x_96; uint8_t x_97; -lean_dec(x_20); +lean_object* x_97; lean_object* x_98; uint8_t x_99; +lean_dec(x_22); lean_dec(x_4); -x_95 = lean_array_get_size(x_18); -x_96 = lean_unsigned_to_nat(2u); -x_97 = lean_nat_dec_eq(x_95, x_96); -lean_dec(x_95); -if (x_97 == 0) +x_97 = lean_array_get_size(x_20); +x_98 = lean_unsigned_to_nat(2u); +x_99 = lean_nat_dec_eq(x_97, x_98); +lean_dec(x_97); +if (x_99 == 0) { -lean_object* x_98; lean_object* x_99; -lean_dec(x_18); +lean_object* x_100; lean_object* x_101; +lean_dec(x_20); +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_3); -x_98 = lean_unsigned_to_nat(0u); -lean_ctor_set(x_14, 1, x_98); -lean_ctor_set(x_14, 0, x_2); -x_99 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_99, 0, x_14); -lean_ctor_set(x_99, 1, x_13); -return x_99; +x_100 = lean_unsigned_to_nat(0u); +lean_ctor_set(x_16, 1, x_100); +lean_ctor_set(x_16, 0, x_2); +x_101 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_101, 0, x_16); +lean_ctor_set(x_101, 1, x_15); +return x_101; } else { -lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; -lean_free_object(x_14); -x_100 = lean_unsigned_to_nat(0u); -x_101 = lean_array_fget(x_18, x_100); -x_102 = lean_unsigned_to_nat(1u); -x_103 = lean_array_fget(x_18, x_102); -lean_dec(x_18); +lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; +lean_free_object(x_16); +x_102 = lean_unsigned_to_nat(0u); +x_103 = lean_array_fget(x_20, x_102); +x_104 = lean_unsigned_to_nat(1u); +x_105 = lean_array_fget(x_20, x_104); +lean_dec(x_20); +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_101); -x_104 = l_Lean_Meta_getLevel(x_101, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_104) == 0) +lean_inc(x_103); +x_106 = l_Lean_Meta_getLevel(x_103, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_106) == 0) { -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_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_box(0); -x_108 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_108, 0, x_105); -lean_ctor_set(x_108, 1, x_107); -x_109 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__10; -x_110 = l_Lean_Expr_const___override(x_109, x_108); -x_111 = l_Lean_mkApp3(x_110, x_101, x_103, x_3); -x_112 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_111, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_106); -return x_112; +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_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_box(0); +x_110 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_110, 0, x_107); +lean_ctor_set(x_110, 1, x_109); +x_111 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__10; +x_112 = l_Lean_Expr_const___override(x_111, x_110); +x_113 = l_Lean_mkApp3(x_112, x_103, x_105, x_3); +x_114 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_113, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_108); +return x_114; } else { -uint8_t x_113; +uint8_t x_115; +lean_dec(x_105); lean_dec(x_103); -lean_dec(x_101); +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_3); lean_dec(x_2); -x_113 = !lean_is_exclusive(x_104); -if (x_113 == 0) +x_115 = !lean_is_exclusive(x_106); +if (x_115 == 0) { -return x_104; +return x_106; } else { -lean_object* x_114; lean_object* x_115; lean_object* x_116; -x_114 = lean_ctor_get(x_104, 0); -x_115 = lean_ctor_get(x_104, 1); -lean_inc(x_115); -lean_inc(x_114); -lean_dec(x_104); -x_116 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_116, 0, x_114); -lean_ctor_set(x_116, 1, x_115); -return x_116; +lean_object* x_116; lean_object* x_117; lean_object* x_118; +x_116 = lean_ctor_get(x_106, 0); +x_117 = lean_ctor_get(x_106, 1); +lean_inc(x_117); +lean_inc(x_116); +lean_dec(x_106); +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; } } } @@ -19720,198 +20011,201 @@ return x_116; } else { -lean_object* x_117; lean_object* x_118; uint8_t x_119; -lean_dec(x_20); +lean_object* x_119; lean_object* x_120; uint8_t x_121; +lean_dec(x_22); lean_dec(x_4); -x_117 = lean_array_get_size(x_18); -x_118 = lean_unsigned_to_nat(2u); -x_119 = lean_nat_dec_eq(x_117, x_118); -lean_dec(x_117); -if (x_119 == 0) +x_119 = lean_array_get_size(x_20); +x_120 = lean_unsigned_to_nat(2u); +x_121 = lean_nat_dec_eq(x_119, x_120); +lean_dec(x_119); +if (x_121 == 0) { -lean_object* x_120; lean_object* x_121; -lean_dec(x_18); +lean_object* x_122; lean_object* x_123; +lean_dec(x_20); +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_3); -x_120 = lean_unsigned_to_nat(0u); -lean_ctor_set(x_14, 1, x_120); -lean_ctor_set(x_14, 0, x_2); -x_121 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_121, 0, x_14); -lean_ctor_set(x_121, 1, x_13); -return x_121; +x_122 = lean_unsigned_to_nat(0u); +lean_ctor_set(x_16, 1, x_122); +lean_ctor_set(x_16, 0, x_2); +x_123 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_123, 0, x_16); +lean_ctor_set(x_123, 1, x_15); +return x_123; } else { -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_free_object(x_14); -x_122 = lean_unsigned_to_nat(0u); -x_123 = lean_array_fget(x_18, x_122); -x_124 = lean_unsigned_to_nat(1u); -x_125 = lean_array_fget(x_18, x_124); -lean_dec(x_18); -x_126 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__13; +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_free_object(x_16); +x_124 = lean_unsigned_to_nat(0u); +x_125 = lean_array_fget(x_20, x_124); +x_126 = lean_unsigned_to_nat(1u); +x_127 = lean_array_fget(x_20, x_126); +lean_dec(x_20); +x_128 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__13; lean_inc(x_3); +lean_inc(x_127); lean_inc(x_125); -lean_inc(x_123); -x_127 = l_Lean_mkApp3(x_126, x_123, x_125, x_3); +x_129 = l_Lean_mkApp3(x_128, x_125, x_127, x_3); +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); lean_inc(x_7); lean_inc(x_6); -x_128 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_127, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_128) == 0) +x_130 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_129, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_130) == 0) { -lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; -x_129 = lean_ctor_get(x_128, 0); -lean_inc(x_129); -x_130 = lean_ctor_get(x_128, 1); -lean_inc(x_130); -lean_dec(x_128); -x_131 = lean_ctor_get(x_129, 0); +lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; +x_131 = lean_ctor_get(x_130, 0); lean_inc(x_131); -x_132 = lean_ctor_get(x_129, 1); +x_132 = lean_ctor_get(x_130, 1); lean_inc(x_132); -lean_dec(x_129); -x_133 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__16; -x_134 = l_Lean_mkApp3(x_133, x_123, x_125, x_3); -x_135 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_131, x_134, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_130); -if (lean_obj_tag(x_135) == 0) -{ -uint8_t x_136; -x_136 = !lean_is_exclusive(x_135); -if (x_136 == 0) +lean_dec(x_130); +x_133 = lean_ctor_get(x_131, 0); +lean_inc(x_133); +x_134 = lean_ctor_get(x_131, 1); +lean_inc(x_134); +lean_dec(x_131); +x_135 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__16; +x_136 = l_Lean_mkApp3(x_135, x_125, x_127, x_3); +x_137 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_133, x_136, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_132); +if (lean_obj_tag(x_137) == 0) { -lean_object* x_137; uint8_t x_138; -x_137 = lean_ctor_get(x_135, 0); +uint8_t x_138; x_138 = !lean_is_exclusive(x_137); if (x_138 == 0) { -lean_object* x_139; lean_object* x_140; -x_139 = lean_ctor_get(x_137, 1); -x_140 = lean_nat_add(x_132, x_139); -lean_dec(x_139); -lean_dec(x_132); -lean_ctor_set(x_137, 1, x_140); -return x_135; +lean_object* x_139; uint8_t x_140; +x_139 = lean_ctor_get(x_137, 0); +x_140 = !lean_is_exclusive(x_139); +if (x_140 == 0) +{ +lean_object* x_141; lean_object* x_142; +x_141 = lean_ctor_get(x_139, 1); +x_142 = lean_nat_add(x_134, x_141); +lean_dec(x_141); +lean_dec(x_134); +lean_ctor_set(x_139, 1, x_142); +return x_137; } else { -lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; -x_141 = lean_ctor_get(x_137, 0); -x_142 = lean_ctor_get(x_137, 1); -lean_inc(x_142); -lean_inc(x_141); -lean_dec(x_137); -x_143 = lean_nat_add(x_132, x_142); -lean_dec(x_142); -lean_dec(x_132); -x_144 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_144, 0, x_141); -lean_ctor_set(x_144, 1, x_143); -lean_ctor_set(x_135, 0, x_144); -return x_135; +lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; +x_143 = lean_ctor_get(x_139, 0); +x_144 = lean_ctor_get(x_139, 1); +lean_inc(x_144); +lean_inc(x_143); +lean_dec(x_139); +x_145 = lean_nat_add(x_134, x_144); +lean_dec(x_144); +lean_dec(x_134); +x_146 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_146, 0, x_143); +lean_ctor_set(x_146, 1, x_145); +lean_ctor_set(x_137, 0, x_146); +return x_137; } } else { -lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; -x_145 = lean_ctor_get(x_135, 0); -x_146 = lean_ctor_get(x_135, 1); -lean_inc(x_146); -lean_inc(x_145); -lean_dec(x_135); -x_147 = lean_ctor_get(x_145, 0); -lean_inc(x_147); -x_148 = lean_ctor_get(x_145, 1); +lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; +x_147 = lean_ctor_get(x_137, 0); +x_148 = lean_ctor_get(x_137, 1); lean_inc(x_148); -if (lean_is_exclusive(x_145)) { - lean_ctor_release(x_145, 0); - lean_ctor_release(x_145, 1); - x_149 = x_145; +lean_inc(x_147); +lean_dec(x_137); +x_149 = lean_ctor_get(x_147, 0); +lean_inc(x_149); +x_150 = lean_ctor_get(x_147, 1); +lean_inc(x_150); +if (lean_is_exclusive(x_147)) { + lean_ctor_release(x_147, 0); + lean_ctor_release(x_147, 1); + x_151 = x_147; } else { - lean_dec_ref(x_145); - x_149 = lean_box(0); + lean_dec_ref(x_147); + x_151 = lean_box(0); } -x_150 = lean_nat_add(x_132, x_148); -lean_dec(x_148); -lean_dec(x_132); -if (lean_is_scalar(x_149)) { - x_151 = lean_alloc_ctor(0, 2, 0); +x_152 = lean_nat_add(x_134, x_150); +lean_dec(x_150); +lean_dec(x_134); +if (lean_is_scalar(x_151)) { + x_153 = lean_alloc_ctor(0, 2, 0); } else { - x_151 = x_149; + x_153 = x_151; } -lean_ctor_set(x_151, 0, x_147); -lean_ctor_set(x_151, 1, 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_146); -return x_152; +lean_ctor_set(x_153, 0, x_149); +lean_ctor_set(x_153, 1, x_152); +x_154 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_154, 0, x_153); +lean_ctor_set(x_154, 1, x_148); +return x_154; } } else { -uint8_t x_153; -lean_dec(x_132); -x_153 = !lean_is_exclusive(x_135); -if (x_153 == 0) +uint8_t x_155; +lean_dec(x_134); +x_155 = !lean_is_exclusive(x_137); +if (x_155 == 0) { -return x_135; +return x_137; } else { -lean_object* x_154; lean_object* x_155; lean_object* x_156; -x_154 = lean_ctor_get(x_135, 0); -x_155 = lean_ctor_get(x_135, 1); -lean_inc(x_155); -lean_inc(x_154); -lean_dec(x_135); -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_156; lean_object* x_157; lean_object* x_158; +x_156 = lean_ctor_get(x_137, 0); +x_157 = lean_ctor_get(x_137, 1); +lean_inc(x_157); +lean_inc(x_156); +lean_dec(x_137); +x_158 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_158, 0, x_156); +lean_ctor_set(x_158, 1, x_157); +return x_158; } } } else { -uint8_t x_157; +uint8_t x_159; +lean_dec(x_127); lean_dec(x_125); -lean_dec(x_123); +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_3); -x_157 = !lean_is_exclusive(x_128); -if (x_157 == 0) +x_159 = !lean_is_exclusive(x_130); +if (x_159 == 0) { -return x_128; +return x_130; } else { -lean_object* x_158; lean_object* x_159; lean_object* x_160; -x_158 = lean_ctor_get(x_128, 0); -x_159 = lean_ctor_get(x_128, 1); -lean_inc(x_159); -lean_inc(x_158); -lean_dec(x_128); -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_160; lean_object* x_161; lean_object* x_162; +x_160 = lean_ctor_get(x_130, 0); +x_161 = lean_ctor_get(x_130, 1); +lean_inc(x_161); +lean_inc(x_160); +lean_dec(x_130); +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; } } } @@ -19919,127 +20213,130 @@ return x_160; } else { -lean_object* x_161; lean_object* x_162; uint8_t x_163; -lean_dec(x_20); +lean_object* x_163; lean_object* x_164; uint8_t x_165; +lean_dec(x_22); lean_dec(x_4); -x_161 = lean_array_get_size(x_18); -x_162 = lean_unsigned_to_nat(1u); -x_163 = lean_nat_dec_eq(x_161, x_162); -lean_dec(x_161); -if (x_163 == 0) +x_163 = lean_array_get_size(x_20); +x_164 = lean_unsigned_to_nat(1u); +x_165 = lean_nat_dec_eq(x_163, x_164); +lean_dec(x_163); +if (x_165 == 0) { -lean_object* x_164; lean_object* x_165; -lean_dec(x_18); +lean_object* x_166; lean_object* x_167; +lean_dec(x_20); +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_3); -x_164 = lean_unsigned_to_nat(0u); -lean_ctor_set(x_14, 1, x_164); -lean_ctor_set(x_14, 0, x_2); -x_165 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_165, 0, x_14); -lean_ctor_set(x_165, 1, x_13); -return x_165; +x_166 = lean_unsigned_to_nat(0u); +lean_ctor_set(x_16, 1, x_166); +lean_ctor_set(x_16, 0, x_2); +x_167 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_167, 0, x_16); +lean_ctor_set(x_167, 1, x_15); +return x_167; } else { -lean_object* x_166; lean_object* x_167; lean_object* x_168; -x_166 = lean_unsigned_to_nat(0u); -x_167 = lean_array_fget(x_18, x_166); -lean_dec(x_18); +lean_object* x_168; lean_object* x_169; lean_object* x_170; +x_168 = lean_unsigned_to_nat(0u); +x_169 = lean_array_fget(x_20, x_168); +lean_dec(x_20); +lean_inc(x_14); +lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -x_168 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot(x_3, x_167, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_168) == 0) +x_170 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot(x_3, x_169, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_170) == 0) { -lean_object* x_169; -x_169 = lean_ctor_get(x_168, 0); -lean_inc(x_169); -if (lean_obj_tag(x_169) == 0) +lean_object* x_171; +x_171 = lean_ctor_get(x_170, 0); +lean_inc(x_171); +if (lean_obj_tag(x_171) == 0) { -uint8_t x_170; +uint8_t x_172; +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); -x_170 = !lean_is_exclusive(x_168); -if (x_170 == 0) +x_172 = !lean_is_exclusive(x_170); +if (x_172 == 0) { -lean_object* x_171; -x_171 = lean_ctor_get(x_168, 0); -lean_dec(x_171); -lean_ctor_set(x_14, 1, x_166); -lean_ctor_set(x_14, 0, x_2); -lean_ctor_set(x_168, 0, x_14); -return x_168; +lean_object* x_173; +x_173 = lean_ctor_get(x_170, 0); +lean_dec(x_173); +lean_ctor_set(x_16, 1, x_168); +lean_ctor_set(x_16, 0, x_2); +lean_ctor_set(x_170, 0, x_16); +return x_170; } else { -lean_object* x_172; lean_object* x_173; -x_172 = lean_ctor_get(x_168, 1); -lean_inc(x_172); -lean_dec(x_168); -lean_ctor_set(x_14, 1, x_166); -lean_ctor_set(x_14, 0, x_2); -x_173 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_173, 0, x_14); -lean_ctor_set(x_173, 1, x_172); -return x_173; +lean_object* x_174; lean_object* x_175; +x_174 = lean_ctor_get(x_170, 1); +lean_inc(x_174); +lean_dec(x_170); +lean_ctor_set(x_16, 1, x_168); +lean_ctor_set(x_16, 0, x_2); +x_175 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_175, 0, x_16); +lean_ctor_set(x_175, 1, x_174); +return x_175; } } else { -lean_object* x_174; lean_object* x_175; lean_object* x_176; -lean_free_object(x_14); -x_174 = lean_ctor_get(x_168, 1); -lean_inc(x_174); -lean_dec(x_168); -x_175 = lean_ctor_get(x_169, 0); -lean_inc(x_175); -lean_dec(x_169); -x_176 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_175, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_174); -return x_176; +lean_object* x_176; lean_object* x_177; lean_object* x_178; +lean_free_object(x_16); +x_176 = lean_ctor_get(x_170, 1); +lean_inc(x_176); +lean_dec(x_170); +x_177 = lean_ctor_get(x_171, 0); +lean_inc(x_177); +lean_dec(x_171); +x_178 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_177, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_176); +return x_178; } } else { -uint8_t x_177; -lean_free_object(x_14); +uint8_t x_179; +lean_free_object(x_16); +lean_dec(x_14); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_2); -x_177 = !lean_is_exclusive(x_168); -if (x_177 == 0) +x_179 = !lean_is_exclusive(x_170); +if (x_179 == 0) { -return x_168; +return x_170; } else { -lean_object* x_178; lean_object* x_179; lean_object* x_180; -x_178 = lean_ctor_get(x_168, 0); -x_179 = lean_ctor_get(x_168, 1); -lean_inc(x_179); -lean_inc(x_178); -lean_dec(x_168); -x_180 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_180, 0, x_178); -lean_ctor_set(x_180, 1, x_179); -return x_180; +lean_object* x_180; lean_object* x_181; lean_object* x_182; +x_180 = lean_ctor_get(x_170, 0); +x_181 = lean_ctor_get(x_170, 1); +lean_inc(x_181); +lean_inc(x_180); +lean_dec(x_170); +x_182 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_182, 0, x_180); +lean_ctor_set(x_182, 1, x_181); +return x_182; } } } @@ -20047,445 +20344,455 @@ return x_180; } else { -lean_object* x_181; lean_object* x_182; uint8_t x_183; -lean_dec(x_20); +lean_object* x_183; lean_object* x_184; uint8_t x_185; +lean_dec(x_22); lean_dec(x_4); -x_181 = lean_array_get_size(x_18); -x_182 = lean_unsigned_to_nat(3u); -x_183 = lean_nat_dec_eq(x_181, x_182); -lean_dec(x_181); -if (x_183 == 0) +x_183 = lean_array_get_size(x_20); +x_184 = lean_unsigned_to_nat(3u); +x_185 = lean_nat_dec_eq(x_183, x_184); +lean_dec(x_183); +if (x_185 == 0) { -lean_object* x_184; lean_object* x_185; -lean_dec(x_18); +lean_object* x_186; lean_object* x_187; +lean_dec(x_20); +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_3); -x_184 = lean_unsigned_to_nat(0u); -lean_ctor_set(x_14, 1, x_184); -lean_ctor_set(x_14, 0, x_2); -x_185 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_185, 0, x_14); -lean_ctor_set(x_185, 1, x_13); -return x_185; +x_186 = lean_unsigned_to_nat(0u); +lean_ctor_set(x_16, 1, x_186); +lean_ctor_set(x_16, 0, x_2); +x_187 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_187, 0, x_16); +lean_ctor_set(x_187, 1, x_15); +return x_187; } else { -lean_object* x_186; lean_object* x_187; -x_186 = lean_unsigned_to_nat(0u); -x_187 = lean_array_fget(x_18, x_186); -if (lean_obj_tag(x_187) == 4) +lean_object* x_188; lean_object* x_189; +x_188 = lean_unsigned_to_nat(0u); +x_189 = lean_array_fget(x_20, x_188); +if (lean_obj_tag(x_189) == 4) { -lean_object* x_188; -x_188 = lean_ctor_get(x_187, 0); -lean_inc(x_188); -if (lean_obj_tag(x_188) == 1) +lean_object* x_190; +x_190 = lean_ctor_get(x_189, 0); +lean_inc(x_190); +if (lean_obj_tag(x_190) == 1) { -lean_object* x_189; -x_189 = lean_ctor_get(x_188, 0); -lean_inc(x_189); -if (lean_obj_tag(x_189) == 0) -{ -lean_object* x_190; lean_object* x_191; lean_object* x_192; uint8_t x_193; -x_190 = lean_ctor_get(x_187, 1); -lean_inc(x_190); -lean_dec(x_187); -x_191 = lean_ctor_get(x_188, 1); +lean_object* x_191; +x_191 = lean_ctor_get(x_190, 0); lean_inc(x_191); -lean_dec(x_188); -x_192 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__1; -x_193 = lean_string_dec_eq(x_191, x_192); -if (x_193 == 0) -{ -lean_object* x_194; uint8_t x_195; -x_194 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__9; -x_195 = lean_string_dec_eq(x_191, x_194); -lean_dec(x_191); -if (x_195 == 0) +if (lean_obj_tag(x_191) == 0) { -lean_object* x_196; +lean_object* x_192; lean_object* x_193; lean_object* x_194; uint8_t x_195; +x_192 = lean_ctor_get(x_189, 1); +lean_inc(x_192); +lean_dec(x_189); +x_193 = lean_ctor_get(x_190, 1); +lean_inc(x_193); lean_dec(x_190); -lean_dec(x_18); +x_194 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__1; +x_195 = lean_string_dec_eq(x_193, x_194); +if (x_195 == 0) +{ +lean_object* x_196; uint8_t x_197; +x_196 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__9; +x_197 = lean_string_dec_eq(x_193, x_196); +lean_dec(x_193); +if (x_197 == 0) +{ +lean_object* x_198; +lean_dec(x_192); +lean_dec(x_20); +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_3); -lean_ctor_set(x_14, 1, x_186); -lean_ctor_set(x_14, 0, x_2); -x_196 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_196, 0, x_14); -lean_ctor_set(x_196, 1, x_13); -return x_196; +lean_ctor_set(x_16, 1, x_188); +lean_ctor_set(x_16, 0, x_2); +x_198 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_198, 0, x_16); +lean_ctor_set(x_198, 1, x_15); +return x_198; } else { -if (lean_obj_tag(x_190) == 0) +if (lean_obj_tag(x_192) == 0) { -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_free_object(x_14); -x_197 = lean_unsigned_to_nat(1u); -x_198 = lean_array_fget(x_18, x_197); -x_199 = lean_unsigned_to_nat(2u); -x_200 = lean_array_fget(x_18, x_199); -lean_dec(x_18); -x_201 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__11___closed__2; -x_202 = l_Lean_mkApp3(x_201, x_198, x_200, x_3); -x_203 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_202, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_203; +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_free_object(x_16); +x_199 = lean_unsigned_to_nat(1u); +x_200 = lean_array_fget(x_20, x_199); +x_201 = lean_unsigned_to_nat(2u); +x_202 = lean_array_fget(x_20, x_201); +lean_dec(x_20); +x_203 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__11___closed__2; +x_204 = l_Lean_mkApp3(x_203, x_200, x_202, x_3); +x_205 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_204, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_205; } else { -lean_object* x_204; -lean_dec(x_190); -lean_dec(x_18); +lean_object* x_206; +lean_dec(x_192); +lean_dec(x_20); +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_3); -lean_ctor_set(x_14, 1, x_186); -lean_ctor_set(x_14, 0, x_2); -x_204 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_204, 0, x_14); -lean_ctor_set(x_204, 1, x_13); -return x_204; +lean_ctor_set(x_16, 1, x_188); +lean_ctor_set(x_16, 0, x_2); +x_206 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_206, 0, x_16); +lean_ctor_set(x_206, 1, x_15); +return x_206; } } } else { -lean_dec(x_191); -if (lean_obj_tag(x_190) == 0) -{ -lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; -lean_free_object(x_14); -x_205 = lean_unsigned_to_nat(1u); -x_206 = lean_array_fget(x_18, x_205); -x_207 = lean_unsigned_to_nat(2u); -x_208 = lean_array_fget(x_18, x_207); -lean_dec(x_18); -x_209 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__12___closed__2; -x_210 = l_Lean_mkApp3(x_209, x_206, x_208, x_3); -x_211 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_210, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_211; +lean_dec(x_193); +if (lean_obj_tag(x_192) == 0) +{ +lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; +lean_free_object(x_16); +x_207 = lean_unsigned_to_nat(1u); +x_208 = lean_array_fget(x_20, x_207); +x_209 = lean_unsigned_to_nat(2u); +x_210 = lean_array_fget(x_20, x_209); +lean_dec(x_20); +x_211 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__12___closed__2; +x_212 = l_Lean_mkApp3(x_211, x_208, x_210, x_3); +x_213 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_212, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_213; } else { -lean_object* x_212; -lean_dec(x_190); -lean_dec(x_18); +lean_object* x_214; +lean_dec(x_192); +lean_dec(x_20); +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_3); -lean_ctor_set(x_14, 1, x_186); -lean_ctor_set(x_14, 0, x_2); -x_212 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_212, 0, x_14); -lean_ctor_set(x_212, 1, x_13); -return x_212; +lean_ctor_set(x_16, 1, x_188); +lean_ctor_set(x_16, 0, x_2); +x_214 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_214, 0, x_16); +lean_ctor_set(x_214, 1, x_15); +return x_214; } } } else { -lean_object* x_213; +lean_object* x_215; +lean_dec(x_191); +lean_dec(x_190); lean_dec(x_189); -lean_dec(x_188); -lean_dec(x_187); -lean_dec(x_18); +lean_dec(x_20); +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_3); -lean_ctor_set(x_14, 1, x_186); -lean_ctor_set(x_14, 0, x_2); -x_213 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_213, 0, x_14); -lean_ctor_set(x_213, 1, x_13); -return x_213; +lean_ctor_set(x_16, 1, x_188); +lean_ctor_set(x_16, 0, x_2); +x_215 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_215, 0, x_16); +lean_ctor_set(x_215, 1, x_15); +return x_215; } } else { -lean_object* x_214; -lean_dec(x_188); -lean_dec(x_187); -lean_dec(x_18); +lean_object* x_216; +lean_dec(x_190); +lean_dec(x_189); +lean_dec(x_20); +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_3); -lean_ctor_set(x_14, 1, x_186); -lean_ctor_set(x_14, 0, x_2); -x_214 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_214, 0, x_14); -lean_ctor_set(x_214, 1, x_13); -return x_214; +lean_ctor_set(x_16, 1, x_188); +lean_ctor_set(x_16, 0, x_2); +x_216 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_216, 0, x_16); +lean_ctor_set(x_216, 1, x_15); +return x_216; } } else { -lean_object* x_215; -lean_dec(x_187); -lean_dec(x_18); +lean_object* x_217; +lean_dec(x_189); +lean_dec(x_20); +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_3); -lean_ctor_set(x_14, 1, x_186); -lean_ctor_set(x_14, 0, x_2); -x_215 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_215, 0, x_14); -lean_ctor_set(x_215, 1, x_13); -return x_215; +lean_ctor_set(x_16, 1, x_188); +lean_ctor_set(x_16, 0, x_2); +x_217 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_217, 0, x_16); +lean_ctor_set(x_217, 1, x_15); +return x_217; } } } } else { -lean_object* x_216; lean_object* x_217; uint8_t x_218; -lean_dec(x_20); +lean_object* x_218; lean_object* x_219; uint8_t x_220; +lean_dec(x_22); lean_dec(x_4); -x_216 = lean_array_get_size(x_18); -x_217 = lean_unsigned_to_nat(3u); -x_218 = lean_nat_dec_eq(x_216, x_217); -lean_dec(x_216); -if (x_218 == 0) +x_218 = lean_array_get_size(x_20); +x_219 = lean_unsigned_to_nat(3u); +x_220 = lean_nat_dec_eq(x_218, x_219); +lean_dec(x_218); +if (x_220 == 0) { -lean_object* x_219; lean_object* x_220; -lean_dec(x_18); +lean_object* x_221; lean_object* x_222; +lean_dec(x_20); +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_3); -x_219 = lean_unsigned_to_nat(0u); -lean_ctor_set(x_14, 1, x_219); -lean_ctor_set(x_14, 0, x_2); -x_220 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_220, 0, x_14); -lean_ctor_set(x_220, 1, x_13); -return x_220; +x_221 = lean_unsigned_to_nat(0u); +lean_ctor_set(x_16, 1, x_221); +lean_ctor_set(x_16, 0, x_2); +x_222 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_222, 0, x_16); +lean_ctor_set(x_222, 1, x_15); +return x_222; } else { -lean_object* x_221; lean_object* x_222; -x_221 = lean_unsigned_to_nat(0u); -x_222 = lean_array_fget(x_18, x_221); -switch (lean_obj_tag(x_222)) { +lean_object* x_223; lean_object* x_224; +x_223 = lean_unsigned_to_nat(0u); +x_224 = lean_array_fget(x_20, x_223); +switch (lean_obj_tag(x_224)) { case 4: { -lean_object* x_223; -x_223 = lean_ctor_get(x_222, 0); -lean_inc(x_223); -if (lean_obj_tag(x_223) == 1) -{ -lean_object* x_224; -x_224 = lean_ctor_get(x_223, 0); -lean_inc(x_224); -if (lean_obj_tag(x_224) == 0) -{ -lean_object* x_225; lean_object* x_226; lean_object* x_227; uint8_t x_228; -x_225 = lean_ctor_get(x_222, 1); +lean_object* x_225; +x_225 = lean_ctor_get(x_224, 0); lean_inc(x_225); -lean_dec(x_222); -x_226 = lean_ctor_get(x_223, 1); +if (lean_obj_tag(x_225) == 1) +{ +lean_object* x_226; +x_226 = lean_ctor_get(x_225, 0); lean_inc(x_226); -lean_dec(x_223); -x_227 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__9; -x_228 = lean_string_dec_eq(x_226, x_227); -if (x_228 == 0) +if (lean_obj_tag(x_226) == 0) { -lean_object* x_229; uint8_t x_230; -x_229 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__1; -x_230 = lean_string_dec_eq(x_226, x_229); -lean_dec(x_226); +lean_object* x_227; lean_object* x_228; lean_object* x_229; uint8_t x_230; +x_227 = lean_ctor_get(x_224, 1); +lean_inc(x_227); +lean_dec(x_224); +x_228 = lean_ctor_get(x_225, 1); +lean_inc(x_228); +lean_dec(x_225); +x_229 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__9; +x_230 = lean_string_dec_eq(x_228, x_229); if (x_230 == 0) { -lean_object* x_231; -lean_dec(x_225); -lean_dec(x_18); +lean_object* x_231; uint8_t x_232; +x_231 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__1; +x_232 = lean_string_dec_eq(x_228, x_231); +lean_dec(x_228); +if (x_232 == 0) +{ +lean_object* x_233; +lean_dec(x_227); +lean_dec(x_20); +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_3); -lean_ctor_set(x_14, 1, x_221); -lean_ctor_set(x_14, 0, x_2); -x_231 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_231, 0, x_14); -lean_ctor_set(x_231, 1, x_13); -return x_231; +lean_ctor_set(x_16, 1, x_223); +lean_ctor_set(x_16, 0, x_2); +x_233 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_233, 0, x_16); +lean_ctor_set(x_233, 1, x_15); +return x_233; } else { -if (lean_obj_tag(x_225) == 0) +if (lean_obj_tag(x_227) == 0) { -lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; -lean_free_object(x_14); -x_232 = lean_unsigned_to_nat(1u); -x_233 = lean_array_fget(x_18, x_232); -x_234 = lean_unsigned_to_nat(2u); -x_235 = lean_array_fget(x_18, x_234); -lean_dec(x_18); -x_236 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__19; -x_237 = l_Lean_mkApp3(x_236, x_233, x_235, x_3); -x_238 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_237, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_238; +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_free_object(x_16); +x_234 = lean_unsigned_to_nat(1u); +x_235 = lean_array_fget(x_20, x_234); +x_236 = lean_unsigned_to_nat(2u); +x_237 = lean_array_fget(x_20, x_236); +lean_dec(x_20); +x_238 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__19; +x_239 = l_Lean_mkApp3(x_238, x_235, x_237, x_3); +x_240 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_239, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_240; } else { -lean_object* x_239; -lean_dec(x_225); -lean_dec(x_18); +lean_object* x_241; +lean_dec(x_227); +lean_dec(x_20); +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_3); -lean_ctor_set(x_14, 1, x_221); -lean_ctor_set(x_14, 0, x_2); -x_239 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_239, 0, x_14); -lean_ctor_set(x_239, 1, x_13); -return x_239; -} +lean_ctor_set(x_16, 1, x_223); +lean_ctor_set(x_16, 0, x_2); +x_241 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_241, 0, x_16); +lean_ctor_set(x_241, 1, x_15); +return x_241; } } -else -{ -lean_dec(x_226); -if (lean_obj_tag(x_225) == 0) -{ -lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; -x_240 = lean_unsigned_to_nat(1u); -x_241 = lean_array_fget(x_18, x_240); -x_242 = lean_unsigned_to_nat(2u); -x_243 = lean_array_fget(x_18, x_242); -lean_dec(x_18); -lean_inc(x_243); -x_244 = l_Lean_Expr_int_x3f(x_243); -if (lean_obj_tag(x_244) == 0) -{ -lean_object* x_245; lean_object* x_246; lean_object* x_247; -lean_free_object(x_14); -x_245 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__22; -x_246 = l_Lean_mkApp3(x_245, x_241, x_243, x_3); -x_247 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_246, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_247; } else { -lean_object* x_248; lean_object* x_249; uint8_t x_250; -x_248 = lean_ctor_get(x_244, 0); -lean_inc(x_248); -lean_dec(x_244); -x_249 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__8; -x_250 = lean_int_dec_eq(x_248, x_249); -lean_dec(x_248); -if (x_250 == 0) +lean_dec(x_228); +if (lean_obj_tag(x_227) == 0) { -lean_object* x_251; lean_object* x_252; lean_object* x_253; -lean_free_object(x_14); -x_251 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__22; -x_252 = l_Lean_mkApp3(x_251, x_241, x_243, x_3); -x_253 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_252, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_253; +lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; +x_242 = lean_unsigned_to_nat(1u); +x_243 = lean_array_fget(x_20, x_242); +x_244 = lean_unsigned_to_nat(2u); +x_245 = lean_array_fget(x_20, x_244); +lean_dec(x_20); +lean_inc(x_245); +x_246 = l_Lean_Expr_int_x3f(x_245); +if (lean_obj_tag(x_246) == 0) +{ +lean_object* x_247; lean_object* x_248; lean_object* x_249; +lean_free_object(x_16); +x_247 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__22; +x_248 = l_Lean_mkApp3(x_247, x_243, x_245, x_3); +x_249 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_248, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_249; +} +else +{ +lean_object* x_250; lean_object* x_251; uint8_t x_252; +x_250 = lean_ctor_get(x_246, 0); +lean_inc(x_250); +lean_dec(x_246); +x_251 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__8; +x_252 = lean_int_dec_eq(x_250, x_251); +lean_dec(x_250); +if (x_252 == 0) +{ +lean_object* x_253; lean_object* x_254; lean_object* x_255; +lean_free_object(x_16); +x_253 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__22; +x_254 = l_Lean_mkApp3(x_253, x_243, x_245, x_3); +x_255 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_254, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_255; } else { -lean_object* x_254; -lean_dec(x_243); -x_254 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality(x_2, x_3, x_241, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_254) == 0) -{ -uint8_t x_255; -x_255 = !lean_is_exclusive(x_254); -if (x_255 == 0) -{ lean_object* x_256; -x_256 = lean_ctor_get(x_254, 0); -lean_ctor_set(x_14, 1, x_240); -lean_ctor_set(x_14, 0, x_256); -lean_ctor_set(x_254, 0, x_14); -return x_254; -} -else +lean_dec(x_245); +x_256 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality(x_2, x_3, x_243, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_256) == 0) +{ +uint8_t x_257; +x_257 = !lean_is_exclusive(x_256); +if (x_257 == 0) { -lean_object* x_257; lean_object* x_258; lean_object* x_259; -x_257 = lean_ctor_get(x_254, 0); -x_258 = lean_ctor_get(x_254, 1); -lean_inc(x_258); -lean_inc(x_257); -lean_dec(x_254); -lean_ctor_set(x_14, 1, x_240); -lean_ctor_set(x_14, 0, x_257); -x_259 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_259, 0, x_14); -lean_ctor_set(x_259, 1, x_258); -return x_259; +lean_object* x_258; +x_258 = lean_ctor_get(x_256, 0); +lean_ctor_set(x_16, 1, x_242); +lean_ctor_set(x_16, 0, x_258); +lean_ctor_set(x_256, 0, x_16); +return x_256; +} +else +{ +lean_object* x_259; lean_object* x_260; lean_object* x_261; +x_259 = lean_ctor_get(x_256, 0); +x_260 = lean_ctor_get(x_256, 1); +lean_inc(x_260); +lean_inc(x_259); +lean_dec(x_256); +lean_ctor_set(x_16, 1, x_242); +lean_ctor_set(x_16, 0, x_259); +x_261 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_261, 0, x_16); +lean_ctor_set(x_261, 1, x_260); +return x_261; } } else { -uint8_t x_260; -lean_free_object(x_14); -x_260 = !lean_is_exclusive(x_254); -if (x_260 == 0) +uint8_t x_262; +lean_free_object(x_16); +x_262 = !lean_is_exclusive(x_256); +if (x_262 == 0) { -return x_254; +return x_256; } else { -lean_object* x_261; lean_object* x_262; lean_object* x_263; -x_261 = lean_ctor_get(x_254, 0); -x_262 = lean_ctor_get(x_254, 1); -lean_inc(x_262); -lean_inc(x_261); -lean_dec(x_254); -x_263 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_263, 0, x_261); -lean_ctor_set(x_263, 1, x_262); -return x_263; +lean_object* x_263; lean_object* x_264; lean_object* x_265; +x_263 = lean_ctor_get(x_256, 0); +x_264 = lean_ctor_get(x_256, 1); +lean_inc(x_264); +lean_inc(x_263); +lean_dec(x_256); +x_265 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_265, 0, x_263); +lean_ctor_set(x_265, 1, x_264); +return x_265; } } } @@ -20493,249 +20800,258 @@ return x_263; } else { -lean_object* x_264; -lean_dec(x_225); -lean_dec(x_18); +lean_object* x_266; +lean_dec(x_227); +lean_dec(x_20); +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_3); -lean_ctor_set(x_14, 1, x_221); -lean_ctor_set(x_14, 0, x_2); -x_264 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_264, 0, x_14); -lean_ctor_set(x_264, 1, x_13); -return x_264; +lean_ctor_set(x_16, 1, x_223); +lean_ctor_set(x_16, 0, x_2); +x_266 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_266, 0, x_16); +lean_ctor_set(x_266, 1, x_15); +return x_266; } } } else { -lean_object* x_265; +lean_object* x_267; +lean_dec(x_226); +lean_dec(x_225); lean_dec(x_224); -lean_dec(x_223); -lean_dec(x_222); -lean_dec(x_18); +lean_dec(x_20); +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_3); -lean_ctor_set(x_14, 1, x_221); -lean_ctor_set(x_14, 0, x_2); -x_265 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_265, 0, x_14); -lean_ctor_set(x_265, 1, x_13); -return x_265; +lean_ctor_set(x_16, 1, x_223); +lean_ctor_set(x_16, 0, x_2); +x_267 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_267, 0, x_16); +lean_ctor_set(x_267, 1, x_15); +return x_267; } } else { -lean_object* x_266; -lean_dec(x_223); -lean_dec(x_222); -lean_dec(x_18); +lean_object* x_268; +lean_dec(x_225); +lean_dec(x_224); +lean_dec(x_20); +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_3); -lean_ctor_set(x_14, 1, x_221); -lean_ctor_set(x_14, 0, x_2); -x_266 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_266, 0, x_14); -lean_ctor_set(x_266, 1, x_13); -return x_266; +lean_ctor_set(x_16, 1, x_223); +lean_ctor_set(x_16, 0, x_2); +x_268 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_268, 0, x_16); +lean_ctor_set(x_268, 1, x_15); +return x_268; } } case 5: { -lean_object* x_267; -x_267 = lean_ctor_get(x_222, 0); -lean_inc(x_267); -if (lean_obj_tag(x_267) == 4) -{ -lean_object* x_268; -x_268 = lean_ctor_get(x_267, 0); -lean_inc(x_268); -if (lean_obj_tag(x_268) == 1) -{ lean_object* x_269; -x_269 = lean_ctor_get(x_268, 0); +x_269 = lean_ctor_get(x_224, 0); lean_inc(x_269); -if (lean_obj_tag(x_269) == 0) +if (lean_obj_tag(x_269) == 4) { -lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; uint8_t x_274; -x_270 = lean_ctor_get(x_222, 1); +lean_object* x_270; +x_270 = lean_ctor_get(x_269, 0); lean_inc(x_270); -lean_dec(x_222); -x_271 = lean_ctor_get(x_267, 1); +if (lean_obj_tag(x_270) == 1) +{ +lean_object* x_271; +x_271 = lean_ctor_get(x_270, 0); lean_inc(x_271); -lean_dec(x_267); -x_272 = lean_ctor_get(x_268, 1); -lean_inc(x_272); -lean_dec(x_268); -x_273 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__4; -x_274 = lean_string_dec_eq(x_272, x_273); -lean_dec(x_272); -if (x_274 == 0) +if (lean_obj_tag(x_271) == 0) { -lean_object* x_275; -lean_dec(x_271); +lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; uint8_t x_276; +x_272 = lean_ctor_get(x_224, 1); +lean_inc(x_272); +lean_dec(x_224); +x_273 = lean_ctor_get(x_269, 1); +lean_inc(x_273); +lean_dec(x_269); +x_274 = lean_ctor_get(x_270, 1); +lean_inc(x_274); lean_dec(x_270); -lean_dec(x_18); +x_275 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__4; +x_276 = lean_string_dec_eq(x_274, x_275); +lean_dec(x_274); +if (x_276 == 0) +{ +lean_object* x_277; +lean_dec(x_273); +lean_dec(x_272); +lean_dec(x_20); +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_3); -lean_ctor_set(x_14, 1, x_221); -lean_ctor_set(x_14, 0, x_2); -x_275 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_275, 0, x_14); -lean_ctor_set(x_275, 1, x_13); -return x_275; +lean_ctor_set(x_16, 1, x_223); +lean_ctor_set(x_16, 0, x_2); +x_277 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_277, 0, x_16); +lean_ctor_set(x_277, 1, x_15); +return x_277; } else { -if (lean_obj_tag(x_271) == 0) +if (lean_obj_tag(x_273) == 0) { -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_free_object(x_14); -x_276 = lean_unsigned_to_nat(1u); -x_277 = lean_array_fget(x_18, x_276); -x_278 = lean_unsigned_to_nat(2u); -x_279 = lean_array_fget(x_18, x_278); -lean_dec(x_18); -x_280 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__25; -x_281 = l_Lean_mkApp4(x_280, x_270, x_277, x_279, x_3); -x_282 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_281, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_282; +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_free_object(x_16); +x_278 = lean_unsigned_to_nat(1u); +x_279 = lean_array_fget(x_20, x_278); +x_280 = lean_unsigned_to_nat(2u); +x_281 = lean_array_fget(x_20, x_280); +lean_dec(x_20); +x_282 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__25; +x_283 = l_Lean_mkApp4(x_282, x_272, x_279, x_281, x_3); +x_284 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_283, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_284; } else { -lean_object* x_283; -lean_dec(x_271); -lean_dec(x_270); -lean_dec(x_18); +lean_object* x_285; +lean_dec(x_273); +lean_dec(x_272); +lean_dec(x_20); +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_3); -lean_ctor_set(x_14, 1, x_221); -lean_ctor_set(x_14, 0, x_2); -x_283 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_283, 0, x_14); -lean_ctor_set(x_283, 1, x_13); -return x_283; +lean_ctor_set(x_16, 1, x_223); +lean_ctor_set(x_16, 0, x_2); +x_285 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_285, 0, x_16); +lean_ctor_set(x_285, 1, x_15); +return x_285; } } } else { -lean_object* x_284; +lean_object* x_286; +lean_dec(x_271); +lean_dec(x_270); lean_dec(x_269); -lean_dec(x_268); -lean_dec(x_267); -lean_dec(x_222); -lean_dec(x_18); +lean_dec(x_224); +lean_dec(x_20); +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_3); -lean_ctor_set(x_14, 1, x_221); -lean_ctor_set(x_14, 0, x_2); -x_284 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_284, 0, x_14); -lean_ctor_set(x_284, 1, x_13); -return x_284; +lean_ctor_set(x_16, 1, x_223); +lean_ctor_set(x_16, 0, x_2); +x_286 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_286, 0, x_16); +lean_ctor_set(x_286, 1, x_15); +return x_286; } } else { -lean_object* x_285; -lean_dec(x_268); -lean_dec(x_267); -lean_dec(x_222); -lean_dec(x_18); +lean_object* x_287; +lean_dec(x_270); +lean_dec(x_269); +lean_dec(x_224); +lean_dec(x_20); +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_3); -lean_ctor_set(x_14, 1, x_221); -lean_ctor_set(x_14, 0, x_2); -x_285 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_285, 0, x_14); -lean_ctor_set(x_285, 1, x_13); -return x_285; +lean_ctor_set(x_16, 1, x_223); +lean_ctor_set(x_16, 0, x_2); +x_287 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_287, 0, x_16); +lean_ctor_set(x_287, 1, x_15); +return x_287; } } else { -lean_object* x_286; -lean_dec(x_267); -lean_dec(x_222); -lean_dec(x_18); +lean_object* x_288; +lean_dec(x_269); +lean_dec(x_224); +lean_dec(x_20); +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_3); -lean_ctor_set(x_14, 1, x_221); -lean_ctor_set(x_14, 0, x_2); -x_286 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_286, 0, x_14); -lean_ctor_set(x_286, 1, x_13); -return x_286; +lean_ctor_set(x_16, 1, x_223); +lean_ctor_set(x_16, 0, x_2); +x_288 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_288, 0, x_16); +lean_ctor_set(x_288, 1, x_15); +return x_288; } } default: { -lean_object* x_287; -lean_dec(x_222); -lean_dec(x_18); +lean_object* x_289; +lean_dec(x_224); +lean_dec(x_20); +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_3); -lean_ctor_set(x_14, 1, x_221); -lean_ctor_set(x_14, 0, x_2); -x_287 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_287, 0, x_14); -lean_ctor_set(x_287, 1, x_13); -return x_287; +lean_ctor_set(x_16, 1, x_223); +lean_ctor_set(x_16, 0, x_2); +x_289 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_289, 0, x_16); +lean_ctor_set(x_289, 1, x_15); +return x_289; } } } @@ -20743,166 +21059,167 @@ return x_287; } else { -lean_object* x_288; lean_object* x_289; lean_object* x_290; uint8_t x_291; -x_288 = lean_ctor_get(x_14, 1); -lean_inc(x_288); -lean_dec(x_14); -x_289 = lean_ctor_get(x_15, 1); -lean_inc(x_289); -lean_dec(x_15); -x_290 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___lambda__2___closed__1; -x_291 = lean_string_dec_eq(x_289, x_290); -if (x_291 == 0) -{ -lean_object* x_292; uint8_t x_293; -x_292 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__1; -x_293 = lean_string_dec_eq(x_289, x_292); +lean_object* x_290; lean_object* x_291; lean_object* x_292; uint8_t x_293; +x_290 = lean_ctor_get(x_16, 1); +lean_inc(x_290); +lean_dec(x_16); +x_291 = lean_ctor_get(x_17, 1); +lean_inc(x_291); +lean_dec(x_17); +x_292 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___lambda__2___closed__1; +x_293 = lean_string_dec_eq(x_291, x_292); if (x_293 == 0) { lean_object* x_294; uint8_t x_295; -x_294 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__1; -x_295 = lean_string_dec_eq(x_289, x_294); +x_294 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__1; +x_295 = lean_string_dec_eq(x_291, x_294); if (x_295 == 0) { lean_object* x_296; uint8_t x_297; -x_296 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__5; -x_297 = lean_string_dec_eq(x_289, x_296); +x_296 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__1; +x_297 = lean_string_dec_eq(x_291, x_296); if (x_297 == 0) { lean_object* x_298; uint8_t x_299; -x_298 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__2; -x_299 = lean_string_dec_eq(x_289, x_298); +x_298 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__5; +x_299 = lean_string_dec_eq(x_291, x_298); if (x_299 == 0) { lean_object* x_300; uint8_t x_301; -x_300 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__3; -x_301 = lean_string_dec_eq(x_289, x_300); +x_300 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__2; +x_301 = lean_string_dec_eq(x_291, x_300); if (x_301 == 0) { lean_object* x_302; uint8_t x_303; -x_302 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__3; -x_303 = lean_string_dec_eq(x_289, x_302); +x_302 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__3; +x_303 = lean_string_dec_eq(x_291, x_302); if (x_303 == 0) { lean_object* x_304; uint8_t x_305; +x_304 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__3; +x_305 = lean_string_dec_eq(x_291, x_304); +if (x_305 == 0) +{ +lean_object* x_306; uint8_t x_307; +lean_dec(x_14); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); -x_304 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__4___closed__1; -x_305 = lean_string_dec_eq(x_289, x_304); -lean_dec(x_289); -if (x_305 == 0) +x_306 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__4___closed__1; +x_307 = lean_string_dec_eq(x_291, x_306); +lean_dec(x_291); +if (x_307 == 0) { -lean_object* x_306; lean_object* x_307; lean_object* x_308; -lean_dec(x_288); +lean_object* x_308; lean_object* x_309; lean_object* x_310; +lean_dec(x_290); lean_dec(x_8); lean_dec(x_4); lean_dec(x_3); -x_306 = lean_unsigned_to_nat(0u); -x_307 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_307, 0, x_2); -lean_ctor_set(x_307, 1, x_306); -x_308 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_308, 0, x_307); -lean_ctor_set(x_308, 1, x_13); -return x_308; +x_308 = lean_unsigned_to_nat(0u); +x_309 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_309, 0, x_2); +lean_ctor_set(x_309, 1, x_308); +x_310 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_310, 0, x_309); +lean_ctor_set(x_310, 1, x_15); +return x_310; } else { -lean_object* x_309; lean_object* x_310; uint8_t x_311; -x_309 = lean_array_get_size(x_288); -lean_dec(x_288); -x_310 = lean_unsigned_to_nat(2u); -x_311 = lean_nat_dec_eq(x_309, x_310); -lean_dec(x_309); -if (x_311 == 0) +lean_object* x_311; lean_object* x_312; uint8_t x_313; +x_311 = lean_array_get_size(x_290); +lean_dec(x_290); +x_312 = lean_unsigned_to_nat(2u); +x_313 = lean_nat_dec_eq(x_311, x_312); +lean_dec(x_311); +if (x_313 == 0) { -lean_object* x_312; lean_object* x_313; lean_object* x_314; +lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_dec(x_8); lean_dec(x_4); lean_dec(x_3); -x_312 = lean_unsigned_to_nat(0u); -x_313 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_313, 0, x_2); -lean_ctor_set(x_313, 1, x_312); -x_314 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_314, 0, x_313); -lean_ctor_set(x_314, 1, x_13); -return x_314; +x_314 = lean_unsigned_to_nat(0u); +x_315 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_315, 0, x_2); +lean_ctor_set(x_315, 1, 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_15); +return x_316; } else { -uint8_t x_315; -x_315 = lean_ctor_get_uint8(x_8, 0); +uint8_t x_317; +x_317 = lean_ctor_get_uint8(x_8, 0); lean_dec(x_8); -if (x_315 == 0) +if (x_317 == 0) { -lean_object* x_316; lean_object* x_317; lean_object* x_318; +lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_dec(x_4); lean_dec(x_3); -x_316 = lean_unsigned_to_nat(0u); -x_317 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_317, 0, x_2); -lean_ctor_set(x_317, 1, x_316); -x_318 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_318, 0, x_317); -lean_ctor_set(x_318, 1, x_13); -return x_318; +x_318 = lean_unsigned_to_nat(0u); +x_319 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_319, 0, x_2); +lean_ctor_set(x_319, 1, x_318); +x_320 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_320, 0, x_319); +lean_ctor_set(x_320, 1, x_15); +return x_320; } else { -lean_object* x_319; lean_object* x_320; uint8_t x_321; -x_319 = lean_ctor_get(x_2, 1); -lean_inc(x_319); -x_320 = lean_ctor_get(x_2, 2); -lean_inc(x_320); -x_321 = l_List_elem___at_Lean_CollectLevelParams_visitExpr___spec__2(x_3, x_320); -if (x_321 == 0) -{ -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_322 = lean_ctor_get(x_2, 3); +lean_object* x_321; lean_object* x_322; uint8_t x_323; +x_321 = lean_ctor_get(x_2, 1); +lean_inc(x_321); +x_322 = lean_ctor_get(x_2, 2); lean_inc(x_322); +x_323 = l_List_elem___at_Lean_CollectLevelParams_visitExpr___spec__2(x_3, x_322); +if (x_323 == 0) +{ +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_324 = lean_ctor_get(x_2, 3); +lean_inc(x_324); lean_dec(x_2); -x_323 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_323, 0, x_3); -lean_ctor_set(x_323, 1, x_320); -x_324 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_324, 0, x_4); -lean_ctor_set(x_324, 1, x_319); -lean_ctor_set(x_324, 2, x_323); -lean_ctor_set(x_324, 3, x_322); -x_325 = lean_unsigned_to_nat(1u); -x_326 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_326, 0, x_324); -lean_ctor_set(x_326, 1, x_325); -x_327 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_327, 0, x_326); -lean_ctor_set(x_327, 1, x_13); -return x_327; +x_325 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_325, 0, x_3); +lean_ctor_set(x_325, 1, x_322); +x_326 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_326, 0, x_4); +lean_ctor_set(x_326, 1, x_321); +lean_ctor_set(x_326, 2, x_325); +lean_ctor_set(x_326, 3, x_324); +x_327 = lean_unsigned_to_nat(1u); +x_328 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_328, 0, x_326); +lean_ctor_set(x_328, 1, x_327); +x_329 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_329, 0, x_328); +lean_ctor_set(x_329, 1, x_15); +return x_329; } else { -lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; +lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_dec(x_3); -x_328 = lean_ctor_get(x_2, 3); -lean_inc(x_328); +x_330 = lean_ctor_get(x_2, 3); +lean_inc(x_330); lean_dec(x_2); -x_329 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_329, 0, x_4); -lean_ctor_set(x_329, 1, x_319); -lean_ctor_set(x_329, 2, x_320); -lean_ctor_set(x_329, 3, x_328); -x_330 = lean_unsigned_to_nat(1u); -x_331 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_331, 0, x_329); -lean_ctor_set(x_331, 1, x_330); -x_332 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_332, 0, x_331); -lean_ctor_set(x_332, 1, x_13); -return x_332; +x_331 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_331, 0, x_4); +lean_ctor_set(x_331, 1, x_321); +lean_ctor_set(x_331, 2, x_322); +lean_ctor_set(x_331, 3, x_330); +x_332 = lean_unsigned_to_nat(1u); +x_333 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_333, 0, x_331); +lean_ctor_set(x_333, 1, x_332); +x_334 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_334, 0, x_333); +lean_ctor_set(x_334, 1, x_15); +return x_334; } } } @@ -20910,1580 +21227,1619 @@ return x_332; } else { -lean_object* x_333; lean_object* x_334; uint8_t x_335; -lean_dec(x_289); +lean_object* x_335; lean_object* x_336; uint8_t x_337; +lean_dec(x_291); lean_dec(x_4); -x_333 = lean_array_get_size(x_288); -x_334 = lean_unsigned_to_nat(2u); -x_335 = lean_nat_dec_eq(x_333, x_334); -lean_dec(x_333); -if (x_335 == 0) +x_335 = lean_array_get_size(x_290); +x_336 = lean_unsigned_to_nat(2u); +x_337 = lean_nat_dec_eq(x_335, x_336); +lean_dec(x_335); +if (x_337 == 0) { -lean_object* x_336; lean_object* x_337; lean_object* x_338; -lean_dec(x_288); +lean_object* x_338; lean_object* x_339; lean_object* x_340; +lean_dec(x_290); +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_3); -x_336 = lean_unsigned_to_nat(0u); -x_337 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_337, 0, x_2); -lean_ctor_set(x_337, 1, x_336); -x_338 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_338, 0, x_337); -lean_ctor_set(x_338, 1, x_13); -return x_338; +x_338 = lean_unsigned_to_nat(0u); +x_339 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_339, 0, x_2); +lean_ctor_set(x_339, 1, x_338); +x_340 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_340, 0, x_339); +lean_ctor_set(x_340, 1, x_15); +return x_340; } else { -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; -x_339 = lean_unsigned_to_nat(0u); -x_340 = lean_array_fget(x_288, x_339); -x_341 = lean_unsigned_to_nat(1u); -x_342 = lean_array_fget(x_288, x_341); -lean_dec(x_288); -x_343 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__1___closed__8; -lean_inc(x_342); -x_344 = l_Lean_Expr_app___override(x_343, x_342); -x_345 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__6; -x_346 = l_Lean_mkApp4(x_345, x_340, x_342, x_344, x_3); -x_347 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_346, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_347; +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; +x_341 = lean_unsigned_to_nat(0u); +x_342 = lean_array_fget(x_290, x_341); +x_343 = lean_unsigned_to_nat(1u); +x_344 = lean_array_fget(x_290, x_343); +lean_dec(x_290); +x_345 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__1___closed__8; +lean_inc(x_344); +x_346 = l_Lean_Expr_app___override(x_345, x_344); +x_347 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__6; +x_348 = l_Lean_mkApp4(x_347, x_342, x_344, x_346, x_3); +x_349 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_348, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_349; } } } else { -lean_object* x_348; lean_object* x_349; uint8_t x_350; -lean_dec(x_289); +lean_object* x_350; lean_object* x_351; uint8_t x_352; +lean_dec(x_291); lean_dec(x_4); -x_348 = lean_array_get_size(x_288); -x_349 = lean_unsigned_to_nat(2u); -x_350 = lean_nat_dec_eq(x_348, x_349); -lean_dec(x_348); -if (x_350 == 0) -{ -lean_object* x_351; lean_object* x_352; lean_object* x_353; -lean_dec(x_288); +x_350 = lean_array_get_size(x_290); +x_351 = lean_unsigned_to_nat(2u); +x_352 = lean_nat_dec_eq(x_350, x_351); +lean_dec(x_350); +if (x_352 == 0) +{ +lean_object* x_353; lean_object* x_354; lean_object* x_355; +lean_dec(x_290); +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_3); -x_351 = lean_unsigned_to_nat(0u); -x_352 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_352, 0, x_2); -lean_ctor_set(x_352, 1, x_351); -x_353 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_353, 0, x_352); -lean_ctor_set(x_353, 1, x_13); -return x_353; +x_353 = lean_unsigned_to_nat(0u); +x_354 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_354, 0, x_2); +lean_ctor_set(x_354, 1, x_353); +x_355 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_355, 0, x_354); +lean_ctor_set(x_355, 1, x_15); +return x_355; } else { -lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; -x_354 = lean_unsigned_to_nat(0u); -x_355 = lean_array_fget(x_288, x_354); -x_356 = lean_unsigned_to_nat(1u); -x_357 = lean_array_fget(x_288, x_356); -lean_dec(x_288); +lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; +x_356 = lean_unsigned_to_nat(0u); +x_357 = lean_array_fget(x_290, x_356); +x_358 = lean_unsigned_to_nat(1u); +x_359 = lean_array_fget(x_290, x_358); +lean_dec(x_290); +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_355); -x_358 = l_Lean_Meta_getLevel(x_355, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_358) == 0) -{ -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; -x_359 = lean_ctor_get(x_358, 0); -lean_inc(x_359); -x_360 = lean_ctor_get(x_358, 1); -lean_inc(x_360); -lean_dec(x_358); -x_361 = lean_box(0); -x_362 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_362, 0, x_359); -lean_ctor_set(x_362, 1, x_361); -x_363 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__8; -x_364 = l_Lean_Expr_const___override(x_363, x_362); -x_365 = l_Lean_mkApp3(x_364, x_355, x_357, x_3); -x_366 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_365, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_360); -return x_366; +lean_inc(x_357); +x_360 = l_Lean_Meta_getLevel(x_357, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_360) == 0) +{ +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; +x_361 = lean_ctor_get(x_360, 0); +lean_inc(x_361); +x_362 = lean_ctor_get(x_360, 1); +lean_inc(x_362); +lean_dec(x_360); +x_363 = lean_box(0); +x_364 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_364, 0, x_361); +lean_ctor_set(x_364, 1, x_363); +x_365 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__8; +x_366 = l_Lean_Expr_const___override(x_365, x_364); +x_367 = l_Lean_mkApp3(x_366, x_357, x_359, x_3); +x_368 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_367, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_362); +return x_368; } else { -lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; +lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; +lean_dec(x_359); lean_dec(x_357); -lean_dec(x_355); +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_3); lean_dec(x_2); -x_367 = lean_ctor_get(x_358, 0); -lean_inc(x_367); -x_368 = lean_ctor_get(x_358, 1); -lean_inc(x_368); -if (lean_is_exclusive(x_358)) { - lean_ctor_release(x_358, 0); - lean_ctor_release(x_358, 1); - x_369 = x_358; +x_369 = lean_ctor_get(x_360, 0); +lean_inc(x_369); +x_370 = lean_ctor_get(x_360, 1); +lean_inc(x_370); +if (lean_is_exclusive(x_360)) { + lean_ctor_release(x_360, 0); + lean_ctor_release(x_360, 1); + x_371 = x_360; } else { - lean_dec_ref(x_358); - x_369 = lean_box(0); + lean_dec_ref(x_360); + x_371 = lean_box(0); } -if (lean_is_scalar(x_369)) { - x_370 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_371)) { + x_372 = lean_alloc_ctor(1, 2, 0); } else { - x_370 = x_369; + x_372 = x_371; } -lean_ctor_set(x_370, 0, x_367); -lean_ctor_set(x_370, 1, x_368); -return x_370; +lean_ctor_set(x_372, 0, x_369); +lean_ctor_set(x_372, 1, x_370); +return x_372; } } } } else { -lean_object* x_371; lean_object* x_372; uint8_t x_373; -lean_dec(x_289); +lean_object* x_373; lean_object* x_374; uint8_t x_375; +lean_dec(x_291); lean_dec(x_4); -x_371 = lean_array_get_size(x_288); -x_372 = lean_unsigned_to_nat(2u); -x_373 = lean_nat_dec_eq(x_371, x_372); -lean_dec(x_371); -if (x_373 == 0) +x_373 = lean_array_get_size(x_290); +x_374 = lean_unsigned_to_nat(2u); +x_375 = lean_nat_dec_eq(x_373, x_374); +lean_dec(x_373); +if (x_375 == 0) { -lean_object* x_374; lean_object* x_375; lean_object* x_376; -lean_dec(x_288); +lean_object* x_376; lean_object* x_377; lean_object* x_378; +lean_dec(x_290); +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_3); -x_374 = lean_unsigned_to_nat(0u); -x_375 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_375, 0, x_2); -lean_ctor_set(x_375, 1, x_374); -x_376 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_376, 0, x_375); -lean_ctor_set(x_376, 1, x_13); -return x_376; +x_376 = lean_unsigned_to_nat(0u); +x_377 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_377, 0, x_2); +lean_ctor_set(x_377, 1, x_376); +x_378 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_378, 0, x_377); +lean_ctor_set(x_378, 1, x_15); +return x_378; } else { -lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; -x_377 = lean_unsigned_to_nat(0u); -x_378 = lean_array_fget(x_288, x_377); -x_379 = lean_unsigned_to_nat(1u); -x_380 = lean_array_fget(x_288, x_379); -lean_dec(x_288); +lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; +x_379 = lean_unsigned_to_nat(0u); +x_380 = lean_array_fget(x_290, x_379); +x_381 = lean_unsigned_to_nat(1u); +x_382 = lean_array_fget(x_290, x_381); +lean_dec(x_290); +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_378); -x_381 = l_Lean_Meta_getLevel(x_378, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_381) == 0) -{ -lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; -x_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 = lean_box(0); -x_385 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_385, 0, x_382); -lean_ctor_set(x_385, 1, x_384); -x_386 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__10; -x_387 = l_Lean_Expr_const___override(x_386, x_385); -x_388 = l_Lean_mkApp3(x_387, x_378, x_380, x_3); -x_389 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_388, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_383); -return x_389; +lean_inc(x_380); +x_383 = l_Lean_Meta_getLevel(x_380, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_383) == 0) +{ +lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; +x_384 = lean_ctor_get(x_383, 0); +lean_inc(x_384); +x_385 = lean_ctor_get(x_383, 1); +lean_inc(x_385); +lean_dec(x_383); +x_386 = lean_box(0); +x_387 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_387, 0, x_384); +lean_ctor_set(x_387, 1, x_386); +x_388 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__10; +x_389 = l_Lean_Expr_const___override(x_388, x_387); +x_390 = l_Lean_mkApp3(x_389, x_380, x_382, x_3); +x_391 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_390, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_385); +return x_391; } else { -lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; +lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; +lean_dec(x_382); lean_dec(x_380); -lean_dec(x_378); +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_3); lean_dec(x_2); -x_390 = lean_ctor_get(x_381, 0); -lean_inc(x_390); -x_391 = lean_ctor_get(x_381, 1); -lean_inc(x_391); -if (lean_is_exclusive(x_381)) { - lean_ctor_release(x_381, 0); - lean_ctor_release(x_381, 1); - x_392 = x_381; +x_392 = lean_ctor_get(x_383, 0); +lean_inc(x_392); +x_393 = lean_ctor_get(x_383, 1); +lean_inc(x_393); +if (lean_is_exclusive(x_383)) { + lean_ctor_release(x_383, 0); + lean_ctor_release(x_383, 1); + x_394 = x_383; } else { - lean_dec_ref(x_381); - x_392 = lean_box(0); + lean_dec_ref(x_383); + x_394 = lean_box(0); } -if (lean_is_scalar(x_392)) { - x_393 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_394)) { + x_395 = lean_alloc_ctor(1, 2, 0); } else { - x_393 = x_392; + x_395 = x_394; } -lean_ctor_set(x_393, 0, x_390); -lean_ctor_set(x_393, 1, x_391); -return x_393; +lean_ctor_set(x_395, 0, x_392); +lean_ctor_set(x_395, 1, x_393); +return x_395; } } } } else { -lean_object* x_394; lean_object* x_395; uint8_t x_396; -lean_dec(x_289); +lean_object* x_396; lean_object* x_397; uint8_t x_398; +lean_dec(x_291); lean_dec(x_4); -x_394 = lean_array_get_size(x_288); -x_395 = lean_unsigned_to_nat(2u); -x_396 = lean_nat_dec_eq(x_394, x_395); -lean_dec(x_394); -if (x_396 == 0) +x_396 = lean_array_get_size(x_290); +x_397 = lean_unsigned_to_nat(2u); +x_398 = lean_nat_dec_eq(x_396, x_397); +lean_dec(x_396); +if (x_398 == 0) { -lean_object* x_397; lean_object* x_398; lean_object* x_399; -lean_dec(x_288); +lean_object* x_399; lean_object* x_400; lean_object* x_401; +lean_dec(x_290); +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_3); -x_397 = lean_unsigned_to_nat(0u); -x_398 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_398, 0, x_2); -lean_ctor_set(x_398, 1, x_397); -x_399 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_399, 0, x_398); -lean_ctor_set(x_399, 1, x_13); -return x_399; +x_399 = lean_unsigned_to_nat(0u); +x_400 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_400, 0, x_2); +lean_ctor_set(x_400, 1, x_399); +x_401 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_401, 0, x_400); +lean_ctor_set(x_401, 1, x_15); +return x_401; } else { -lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; -x_400 = lean_unsigned_to_nat(0u); -x_401 = lean_array_fget(x_288, x_400); -x_402 = lean_unsigned_to_nat(1u); -x_403 = lean_array_fget(x_288, x_402); -lean_dec(x_288); -x_404 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__13; +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; +x_402 = lean_unsigned_to_nat(0u); +x_403 = lean_array_fget(x_290, x_402); +x_404 = lean_unsigned_to_nat(1u); +x_405 = lean_array_fget(x_290, x_404); +lean_dec(x_290); +x_406 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__13; lean_inc(x_3); +lean_inc(x_405); lean_inc(x_403); -lean_inc(x_401); -x_405 = l_Lean_mkApp3(x_404, x_401, x_403, x_3); +x_407 = l_Lean_mkApp3(x_406, x_403, x_405, x_3); +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); lean_inc(x_7); lean_inc(x_6); -x_406 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_405, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_406) == 0) +x_408 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_407, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_408) == 0) { -lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; -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_ctor_get(x_407, 0); +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; +x_409 = lean_ctor_get(x_408, 0); lean_inc(x_409); -x_410 = lean_ctor_get(x_407, 1); +x_410 = lean_ctor_get(x_408, 1); lean_inc(x_410); -lean_dec(x_407); -x_411 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__16; -x_412 = l_Lean_mkApp3(x_411, x_401, x_403, x_3); -x_413 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_409, x_412, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_408); -if (lean_obj_tag(x_413) == 0) -{ -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; -x_414 = lean_ctor_get(x_413, 0); -lean_inc(x_414); -x_415 = lean_ctor_get(x_413, 1); -lean_inc(x_415); -if (lean_is_exclusive(x_413)) { - lean_ctor_release(x_413, 0); - lean_ctor_release(x_413, 1); - x_416 = x_413; -} else { - lean_dec_ref(x_413); - x_416 = lean_box(0); -} -x_417 = lean_ctor_get(x_414, 0); +lean_dec(x_408); +x_411 = lean_ctor_get(x_409, 0); +lean_inc(x_411); +x_412 = lean_ctor_get(x_409, 1); +lean_inc(x_412); +lean_dec(x_409); +x_413 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__16; +x_414 = l_Lean_mkApp3(x_413, x_403, x_405, x_3); +x_415 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_411, x_414, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_410); +if (lean_obj_tag(x_415) == 0) +{ +lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; +x_416 = lean_ctor_get(x_415, 0); +lean_inc(x_416); +x_417 = lean_ctor_get(x_415, 1); lean_inc(x_417); -x_418 = lean_ctor_get(x_414, 1); -lean_inc(x_418); -if (lean_is_exclusive(x_414)) { - lean_ctor_release(x_414, 0); - lean_ctor_release(x_414, 1); - x_419 = x_414; +if (lean_is_exclusive(x_415)) { + lean_ctor_release(x_415, 0); + lean_ctor_release(x_415, 1); + x_418 = x_415; } else { - lean_dec_ref(x_414); - x_419 = lean_box(0); + lean_dec_ref(x_415); + x_418 = lean_box(0); +} +x_419 = lean_ctor_get(x_416, 0); +lean_inc(x_419); +x_420 = lean_ctor_get(x_416, 1); +lean_inc(x_420); +if (lean_is_exclusive(x_416)) { + lean_ctor_release(x_416, 0); + lean_ctor_release(x_416, 1); + x_421 = x_416; +} else { + lean_dec_ref(x_416); + x_421 = lean_box(0); } -x_420 = lean_nat_add(x_410, x_418); -lean_dec(x_418); -lean_dec(x_410); -if (lean_is_scalar(x_419)) { - x_421 = lean_alloc_ctor(0, 2, 0); +x_422 = lean_nat_add(x_412, x_420); +lean_dec(x_420); +lean_dec(x_412); +if (lean_is_scalar(x_421)) { + x_423 = lean_alloc_ctor(0, 2, 0); } else { - x_421 = x_419; + x_423 = x_421; } -lean_ctor_set(x_421, 0, x_417); -lean_ctor_set(x_421, 1, x_420); -if (lean_is_scalar(x_416)) { - x_422 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_423, 0, x_419); +lean_ctor_set(x_423, 1, x_422); +if (lean_is_scalar(x_418)) { + x_424 = lean_alloc_ctor(0, 2, 0); } else { - x_422 = x_416; + x_424 = x_418; } -lean_ctor_set(x_422, 0, x_421); -lean_ctor_set(x_422, 1, x_415); -return x_422; +lean_ctor_set(x_424, 0, x_423); +lean_ctor_set(x_424, 1, x_417); +return x_424; } else { -lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; -lean_dec(x_410); -x_423 = lean_ctor_get(x_413, 0); -lean_inc(x_423); -x_424 = lean_ctor_get(x_413, 1); -lean_inc(x_424); -if (lean_is_exclusive(x_413)) { - lean_ctor_release(x_413, 0); - lean_ctor_release(x_413, 1); - x_425 = x_413; +lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; +lean_dec(x_412); +x_425 = lean_ctor_get(x_415, 0); +lean_inc(x_425); +x_426 = lean_ctor_get(x_415, 1); +lean_inc(x_426); +if (lean_is_exclusive(x_415)) { + lean_ctor_release(x_415, 0); + lean_ctor_release(x_415, 1); + x_427 = x_415; } else { - lean_dec_ref(x_413); - x_425 = lean_box(0); + lean_dec_ref(x_415); + x_427 = lean_box(0); } -if (lean_is_scalar(x_425)) { - x_426 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_427)) { + x_428 = lean_alloc_ctor(1, 2, 0); } else { - x_426 = x_425; + x_428 = x_427; } -lean_ctor_set(x_426, 0, x_423); -lean_ctor_set(x_426, 1, x_424); -return x_426; +lean_ctor_set(x_428, 0, x_425); +lean_ctor_set(x_428, 1, x_426); +return x_428; } } else { -lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; +lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; +lean_dec(x_405); lean_dec(x_403); -lean_dec(x_401); +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_3); -x_427 = lean_ctor_get(x_406, 0); -lean_inc(x_427); -x_428 = lean_ctor_get(x_406, 1); -lean_inc(x_428); -if (lean_is_exclusive(x_406)) { - lean_ctor_release(x_406, 0); - lean_ctor_release(x_406, 1); - x_429 = x_406; +x_429 = lean_ctor_get(x_408, 0); +lean_inc(x_429); +x_430 = lean_ctor_get(x_408, 1); +lean_inc(x_430); +if (lean_is_exclusive(x_408)) { + lean_ctor_release(x_408, 0); + lean_ctor_release(x_408, 1); + x_431 = x_408; } else { - lean_dec_ref(x_406); - x_429 = lean_box(0); + lean_dec_ref(x_408); + x_431 = lean_box(0); } -if (lean_is_scalar(x_429)) { - x_430 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_431)) { + x_432 = lean_alloc_ctor(1, 2, 0); } else { - x_430 = x_429; + x_432 = x_431; } -lean_ctor_set(x_430, 0, x_427); -lean_ctor_set(x_430, 1, x_428); -return x_430; +lean_ctor_set(x_432, 0, x_429); +lean_ctor_set(x_432, 1, x_430); +return x_432; } } } } else { -lean_object* x_431; lean_object* x_432; uint8_t x_433; -lean_dec(x_289); +lean_object* x_433; lean_object* x_434; uint8_t x_435; +lean_dec(x_291); lean_dec(x_4); -x_431 = lean_array_get_size(x_288); -x_432 = lean_unsigned_to_nat(1u); -x_433 = lean_nat_dec_eq(x_431, x_432); -lean_dec(x_431); -if (x_433 == 0) +x_433 = lean_array_get_size(x_290); +x_434 = lean_unsigned_to_nat(1u); +x_435 = lean_nat_dec_eq(x_433, x_434); +lean_dec(x_433); +if (x_435 == 0) { -lean_object* x_434; lean_object* x_435; lean_object* x_436; -lean_dec(x_288); +lean_object* x_436; lean_object* x_437; lean_object* x_438; +lean_dec(x_290); +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_3); -x_434 = lean_unsigned_to_nat(0u); -x_435 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_435, 0, x_2); -lean_ctor_set(x_435, 1, x_434); -x_436 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_436, 0, x_435); -lean_ctor_set(x_436, 1, x_13); -return x_436; +x_436 = lean_unsigned_to_nat(0u); +x_437 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_437, 0, x_2); +lean_ctor_set(x_437, 1, x_436); +x_438 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_438, 0, x_437); +lean_ctor_set(x_438, 1, x_15); +return x_438; } else { -lean_object* x_437; lean_object* x_438; lean_object* x_439; -x_437 = lean_unsigned_to_nat(0u); -x_438 = lean_array_fget(x_288, x_437); -lean_dec(x_288); +lean_object* x_439; lean_object* x_440; lean_object* x_441; +x_439 = lean_unsigned_to_nat(0u); +x_440 = lean_array_fget(x_290, x_439); +lean_dec(x_290); +lean_inc(x_14); +lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -x_439 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot(x_3, x_438, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_439) == 0) +x_441 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot(x_3, x_440, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_441) == 0) { -lean_object* x_440; -x_440 = lean_ctor_get(x_439, 0); -lean_inc(x_440); -if (lean_obj_tag(x_440) == 0) +lean_object* x_442; +x_442 = lean_ctor_get(x_441, 0); +lean_inc(x_442); +if (lean_obj_tag(x_442) == 0) { -lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; +lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; +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); -x_441 = lean_ctor_get(x_439, 1); -lean_inc(x_441); -if (lean_is_exclusive(x_439)) { - lean_ctor_release(x_439, 0); - lean_ctor_release(x_439, 1); - x_442 = x_439; +x_443 = lean_ctor_get(x_441, 1); +lean_inc(x_443); +if (lean_is_exclusive(x_441)) { + lean_ctor_release(x_441, 0); + lean_ctor_release(x_441, 1); + x_444 = x_441; } else { - lean_dec_ref(x_439); - x_442 = lean_box(0); -} -x_443 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_443, 0, x_2); -lean_ctor_set(x_443, 1, x_437); -if (lean_is_scalar(x_442)) { - x_444 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_441); + x_444 = lean_box(0); +} +x_445 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_445, 0, x_2); +lean_ctor_set(x_445, 1, x_439); +if (lean_is_scalar(x_444)) { + x_446 = lean_alloc_ctor(0, 2, 0); } else { - x_444 = x_442; + x_446 = x_444; } -lean_ctor_set(x_444, 0, x_443); -lean_ctor_set(x_444, 1, x_441); -return x_444; +lean_ctor_set(x_446, 0, x_445); +lean_ctor_set(x_446, 1, x_443); +return x_446; } else { -lean_object* x_445; lean_object* x_446; lean_object* x_447; -x_445 = lean_ctor_get(x_439, 1); -lean_inc(x_445); -lean_dec(x_439); -x_446 = lean_ctor_get(x_440, 0); -lean_inc(x_446); -lean_dec(x_440); -x_447 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_446, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_445); -return x_447; +lean_object* x_447; lean_object* x_448; lean_object* x_449; +x_447 = lean_ctor_get(x_441, 1); +lean_inc(x_447); +lean_dec(x_441); +x_448 = lean_ctor_get(x_442, 0); +lean_inc(x_448); +lean_dec(x_442); +x_449 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_448, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_447); +return x_449; } } else { -lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; +lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; +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_2); -x_448 = lean_ctor_get(x_439, 0); -lean_inc(x_448); -x_449 = lean_ctor_get(x_439, 1); -lean_inc(x_449); -if (lean_is_exclusive(x_439)) { - lean_ctor_release(x_439, 0); - lean_ctor_release(x_439, 1); - x_450 = x_439; +x_450 = lean_ctor_get(x_441, 0); +lean_inc(x_450); +x_451 = lean_ctor_get(x_441, 1); +lean_inc(x_451); +if (lean_is_exclusive(x_441)) { + lean_ctor_release(x_441, 0); + lean_ctor_release(x_441, 1); + x_452 = x_441; } else { - lean_dec_ref(x_439); - x_450 = lean_box(0); + lean_dec_ref(x_441); + x_452 = lean_box(0); } -if (lean_is_scalar(x_450)) { - x_451 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_452)) { + x_453 = lean_alloc_ctor(1, 2, 0); } else { - x_451 = x_450; + x_453 = x_452; } -lean_ctor_set(x_451, 0, x_448); -lean_ctor_set(x_451, 1, x_449); -return x_451; +lean_ctor_set(x_453, 0, x_450); +lean_ctor_set(x_453, 1, x_451); +return x_453; } } } } else { -lean_object* x_452; lean_object* x_453; uint8_t x_454; -lean_dec(x_289); +lean_object* x_454; lean_object* x_455; uint8_t x_456; +lean_dec(x_291); lean_dec(x_4); -x_452 = lean_array_get_size(x_288); -x_453 = lean_unsigned_to_nat(3u); -x_454 = lean_nat_dec_eq(x_452, x_453); -lean_dec(x_452); -if (x_454 == 0) -{ -lean_object* x_455; lean_object* x_456; lean_object* x_457; -lean_dec(x_288); +x_454 = lean_array_get_size(x_290); +x_455 = lean_unsigned_to_nat(3u); +x_456 = lean_nat_dec_eq(x_454, x_455); +lean_dec(x_454); +if (x_456 == 0) +{ +lean_object* x_457; lean_object* x_458; lean_object* x_459; +lean_dec(x_290); +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_3); -x_455 = lean_unsigned_to_nat(0u); -x_456 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_456, 0, x_2); -lean_ctor_set(x_456, 1, x_455); -x_457 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_457, 0, x_456); -lean_ctor_set(x_457, 1, x_13); -return x_457; +x_457 = lean_unsigned_to_nat(0u); +x_458 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_458, 0, x_2); +lean_ctor_set(x_458, 1, x_457); +x_459 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_459, 0, x_458); +lean_ctor_set(x_459, 1, x_15); +return x_459; } else { -lean_object* x_458; lean_object* x_459; -x_458 = lean_unsigned_to_nat(0u); -x_459 = lean_array_fget(x_288, x_458); -if (lean_obj_tag(x_459) == 4) -{ -lean_object* x_460; -x_460 = lean_ctor_get(x_459, 0); -lean_inc(x_460); -if (lean_obj_tag(x_460) == 1) -{ -lean_object* x_461; -x_461 = lean_ctor_get(x_460, 0); -lean_inc(x_461); -if (lean_obj_tag(x_461) == 0) +lean_object* x_460; lean_object* x_461; +x_460 = lean_unsigned_to_nat(0u); +x_461 = lean_array_fget(x_290, x_460); +if (lean_obj_tag(x_461) == 4) { -lean_object* x_462; lean_object* x_463; lean_object* x_464; uint8_t x_465; -x_462 = lean_ctor_get(x_459, 1); +lean_object* x_462; +x_462 = lean_ctor_get(x_461, 0); lean_inc(x_462); -lean_dec(x_459); -x_463 = lean_ctor_get(x_460, 1); +if (lean_obj_tag(x_462) == 1) +{ +lean_object* x_463; +x_463 = lean_ctor_get(x_462, 0); lean_inc(x_463); -lean_dec(x_460); -x_464 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__1; -x_465 = lean_string_dec_eq(x_463, x_464); -if (x_465 == 0) -{ -lean_object* x_466; uint8_t x_467; -x_466 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__9; -x_467 = lean_string_dec_eq(x_463, x_466); -lean_dec(x_463); -if (x_467 == 0) +if (lean_obj_tag(x_463) == 0) { -lean_object* x_468; lean_object* x_469; +lean_object* x_464; lean_object* x_465; lean_object* x_466; uint8_t x_467; +x_464 = lean_ctor_get(x_461, 1); +lean_inc(x_464); +lean_dec(x_461); +x_465 = lean_ctor_get(x_462, 1); +lean_inc(x_465); lean_dec(x_462); -lean_dec(x_288); +x_466 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__1; +x_467 = lean_string_dec_eq(x_465, x_466); +if (x_467 == 0) +{ +lean_object* x_468; uint8_t x_469; +x_468 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__9; +x_469 = lean_string_dec_eq(x_465, x_468); +lean_dec(x_465); +if (x_469 == 0) +{ +lean_object* x_470; lean_object* x_471; +lean_dec(x_464); +lean_dec(x_290); +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_3); -x_468 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_468, 0, x_2); -lean_ctor_set(x_468, 1, x_458); -x_469 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_469, 0, x_468); -lean_ctor_set(x_469, 1, x_13); -return x_469; +x_470 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_470, 0, x_2); +lean_ctor_set(x_470, 1, x_460); +x_471 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_471, 0, x_470); +lean_ctor_set(x_471, 1, x_15); +return x_471; } else { -if (lean_obj_tag(x_462) == 0) +if (lean_obj_tag(x_464) == 0) { -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; -x_470 = lean_unsigned_to_nat(1u); -x_471 = lean_array_fget(x_288, x_470); -x_472 = lean_unsigned_to_nat(2u); -x_473 = lean_array_fget(x_288, x_472); -lean_dec(x_288); -x_474 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__11___closed__2; -x_475 = l_Lean_mkApp3(x_474, x_471, x_473, x_3); -x_476 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_475, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_476; +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; +x_472 = lean_unsigned_to_nat(1u); +x_473 = lean_array_fget(x_290, x_472); +x_474 = lean_unsigned_to_nat(2u); +x_475 = lean_array_fget(x_290, x_474); +lean_dec(x_290); +x_476 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__11___closed__2; +x_477 = l_Lean_mkApp3(x_476, x_473, x_475, x_3); +x_478 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_477, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_478; } else { -lean_object* x_477; lean_object* x_478; -lean_dec(x_462); -lean_dec(x_288); +lean_object* x_479; lean_object* x_480; +lean_dec(x_464); +lean_dec(x_290); +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_3); -x_477 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_477, 0, x_2); -lean_ctor_set(x_477, 1, x_458); -x_478 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_478, 0, x_477); -lean_ctor_set(x_478, 1, x_13); -return x_478; +x_479 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_479, 0, x_2); +lean_ctor_set(x_479, 1, x_460); +x_480 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_480, 0, x_479); +lean_ctor_set(x_480, 1, x_15); +return x_480; } } } else { -lean_dec(x_463); -if (lean_obj_tag(x_462) == 0) -{ -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; -x_479 = lean_unsigned_to_nat(1u); -x_480 = lean_array_fget(x_288, x_479); -x_481 = lean_unsigned_to_nat(2u); -x_482 = lean_array_fget(x_288, x_481); -lean_dec(x_288); -x_483 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__12___closed__2; -x_484 = l_Lean_mkApp3(x_483, x_480, x_482, x_3); -x_485 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_484, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_485; -} -else +lean_dec(x_465); +if (lean_obj_tag(x_464) == 0) { -lean_object* x_486; lean_object* x_487; -lean_dec(x_462); -lean_dec(x_288); -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_3); -x_486 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_486, 0, x_2); -lean_ctor_set(x_486, 1, x_458); -x_487 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_487, 0, x_486); -lean_ctor_set(x_487, 1, x_13); +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; +x_481 = lean_unsigned_to_nat(1u); +x_482 = lean_array_fget(x_290, x_481); +x_483 = lean_unsigned_to_nat(2u); +x_484 = lean_array_fget(x_290, x_483); +lean_dec(x_290); +x_485 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__12___closed__2; +x_486 = l_Lean_mkApp3(x_485, x_482, x_484, x_3); +x_487 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_486, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); return x_487; } -} -} else { lean_object* x_488; lean_object* x_489; -lean_dec(x_461); -lean_dec(x_460); -lean_dec(x_459); -lean_dec(x_288); +lean_dec(x_464); +lean_dec(x_290); +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_3); x_488 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_488, 0, x_2); -lean_ctor_set(x_488, 1, x_458); +lean_ctor_set(x_488, 1, x_460); x_489 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_489, 0, x_488); -lean_ctor_set(x_489, 1, x_13); +lean_ctor_set(x_489, 1, x_15); return x_489; } } +} else { lean_object* x_490; lean_object* x_491; -lean_dec(x_460); -lean_dec(x_459); -lean_dec(x_288); +lean_dec(x_463); +lean_dec(x_462); +lean_dec(x_461); +lean_dec(x_290); +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_3); x_490 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_490, 0, x_2); -lean_ctor_set(x_490, 1, x_458); +lean_ctor_set(x_490, 1, x_460); x_491 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_491, 0, x_490); -lean_ctor_set(x_491, 1, x_13); +lean_ctor_set(x_491, 1, x_15); return x_491; } } else { lean_object* x_492; lean_object* x_493; -lean_dec(x_459); -lean_dec(x_288); +lean_dec(x_462); +lean_dec(x_461); +lean_dec(x_290); +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_3); x_492 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_492, 0, x_2); -lean_ctor_set(x_492, 1, x_458); +lean_ctor_set(x_492, 1, x_460); x_493 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_493, 0, x_492); -lean_ctor_set(x_493, 1, x_13); +lean_ctor_set(x_493, 1, x_15); return x_493; } } -} -} else { -lean_object* x_494; lean_object* x_495; uint8_t x_496; -lean_dec(x_289); -lean_dec(x_4); -x_494 = lean_array_get_size(x_288); -x_495 = lean_unsigned_to_nat(3u); -x_496 = lean_nat_dec_eq(x_494, x_495); -lean_dec(x_494); -if (x_496 == 0) -{ -lean_object* x_497; lean_object* x_498; lean_object* x_499; -lean_dec(x_288); +lean_object* x_494; lean_object* x_495; +lean_dec(x_461); +lean_dec(x_290); +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_3); -x_497 = lean_unsigned_to_nat(0u); -x_498 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_498, 0, x_2); -lean_ctor_set(x_498, 1, 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_13); -return x_499; +x_494 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_494, 0, x_2); +lean_ctor_set(x_494, 1, x_460); +x_495 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_495, 0, x_494); +lean_ctor_set(x_495, 1, x_15); +return x_495; +} +} +} } else { -lean_object* x_500; lean_object* x_501; -x_500 = lean_unsigned_to_nat(0u); -x_501 = lean_array_fget(x_288, x_500); -switch (lean_obj_tag(x_501)) { -case 4: +lean_object* x_496; lean_object* x_497; uint8_t x_498; +lean_dec(x_291); +lean_dec(x_4); +x_496 = lean_array_get_size(x_290); +x_497 = lean_unsigned_to_nat(3u); +x_498 = lean_nat_dec_eq(x_496, x_497); +lean_dec(x_496); +if (x_498 == 0) { -lean_object* x_502; -x_502 = lean_ctor_get(x_501, 0); -lean_inc(x_502); -if (lean_obj_tag(x_502) == 1) +lean_object* x_499; lean_object* x_500; lean_object* x_501; +lean_dec(x_290); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_3); +x_499 = lean_unsigned_to_nat(0u); +x_500 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_500, 0, x_2); +lean_ctor_set(x_500, 1, x_499); +x_501 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_501, 0, x_500); +lean_ctor_set(x_501, 1, x_15); +return x_501; +} +else { -lean_object* x_503; -x_503 = lean_ctor_get(x_502, 0); -lean_inc(x_503); -if (lean_obj_tag(x_503) == 0) +lean_object* x_502; lean_object* x_503; +x_502 = lean_unsigned_to_nat(0u); +x_503 = lean_array_fget(x_290, x_502); +switch (lean_obj_tag(x_503)) { +case 4: { -lean_object* x_504; lean_object* x_505; lean_object* x_506; uint8_t x_507; -x_504 = lean_ctor_get(x_501, 1); +lean_object* x_504; +x_504 = lean_ctor_get(x_503, 0); lean_inc(x_504); -lean_dec(x_501); -x_505 = lean_ctor_get(x_502, 1); +if (lean_obj_tag(x_504) == 1) +{ +lean_object* x_505; +x_505 = lean_ctor_get(x_504, 0); lean_inc(x_505); -lean_dec(x_502); -x_506 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__9; -x_507 = lean_string_dec_eq(x_505, x_506); -if (x_507 == 0) -{ -lean_object* x_508; uint8_t x_509; -x_508 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__1; -x_509 = lean_string_dec_eq(x_505, x_508); -lean_dec(x_505); -if (x_509 == 0) +if (lean_obj_tag(x_505) == 0) { -lean_object* x_510; lean_object* x_511; +lean_object* x_506; lean_object* x_507; lean_object* x_508; uint8_t x_509; +x_506 = lean_ctor_get(x_503, 1); +lean_inc(x_506); +lean_dec(x_503); +x_507 = lean_ctor_get(x_504, 1); +lean_inc(x_507); lean_dec(x_504); -lean_dec(x_288); +x_508 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__9; +x_509 = lean_string_dec_eq(x_507, x_508); +if (x_509 == 0) +{ +lean_object* x_510; uint8_t x_511; +x_510 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__1; +x_511 = lean_string_dec_eq(x_507, x_510); +lean_dec(x_507); +if (x_511 == 0) +{ +lean_object* x_512; lean_object* x_513; +lean_dec(x_506); +lean_dec(x_290); +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_3); -x_510 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_510, 0, x_2); -lean_ctor_set(x_510, 1, x_500); -x_511 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_511, 0, x_510); -lean_ctor_set(x_511, 1, x_13); -return x_511; +x_512 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_512, 0, x_2); +lean_ctor_set(x_512, 1, x_502); +x_513 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_513, 0, x_512); +lean_ctor_set(x_513, 1, x_15); +return x_513; } else { -if (lean_obj_tag(x_504) == 0) +if (lean_obj_tag(x_506) == 0) { -lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; -x_512 = lean_unsigned_to_nat(1u); -x_513 = lean_array_fget(x_288, x_512); -x_514 = lean_unsigned_to_nat(2u); -x_515 = lean_array_fget(x_288, x_514); -lean_dec(x_288); -x_516 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__19; -x_517 = l_Lean_mkApp3(x_516, x_513, x_515, x_3); -x_518 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_517, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_518; +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; +x_514 = lean_unsigned_to_nat(1u); +x_515 = lean_array_fget(x_290, x_514); +x_516 = lean_unsigned_to_nat(2u); +x_517 = lean_array_fget(x_290, x_516); +lean_dec(x_290); +x_518 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__19; +x_519 = l_Lean_mkApp3(x_518, x_515, x_517, x_3); +x_520 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_519, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_520; } else { -lean_object* x_519; lean_object* x_520; -lean_dec(x_504); -lean_dec(x_288); +lean_object* x_521; lean_object* x_522; +lean_dec(x_506); +lean_dec(x_290); +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_3); -x_519 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_519, 0, x_2); -lean_ctor_set(x_519, 1, x_500); -x_520 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_520, 0, x_519); -lean_ctor_set(x_520, 1, x_13); -return x_520; -} +x_521 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_521, 0, x_2); +lean_ctor_set(x_521, 1, x_502); +x_522 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_522, 0, x_521); +lean_ctor_set(x_522, 1, x_15); +return x_522; } } -else -{ -lean_dec(x_505); -if (lean_obj_tag(x_504) == 0) -{ -lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; -x_521 = lean_unsigned_to_nat(1u); -x_522 = lean_array_fget(x_288, x_521); -x_523 = lean_unsigned_to_nat(2u); -x_524 = lean_array_fget(x_288, x_523); -lean_dec(x_288); -lean_inc(x_524); -x_525 = l_Lean_Expr_int_x3f(x_524); -if (lean_obj_tag(x_525) == 0) -{ -lean_object* x_526; lean_object* x_527; lean_object* x_528; -x_526 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__22; -x_527 = l_Lean_mkApp3(x_526, x_522, x_524, x_3); -x_528 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_527, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_528; -} -else -{ -lean_object* x_529; lean_object* x_530; uint8_t x_531; -x_529 = lean_ctor_get(x_525, 0); -lean_inc(x_529); -lean_dec(x_525); -x_530 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__8; -x_531 = lean_int_dec_eq(x_529, x_530); -lean_dec(x_529); -if (x_531 == 0) -{ -lean_object* x_532; lean_object* x_533; lean_object* x_534; -x_532 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__22; -x_533 = l_Lean_mkApp3(x_532, x_522, x_524, x_3); -x_534 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_533, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_534; } else { -lean_object* x_535; -lean_dec(x_524); -x_535 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality(x_2, x_3, x_522, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_535) == 0) +lean_dec(x_507); +if (lean_obj_tag(x_506) == 0) { -lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; -x_536 = lean_ctor_get(x_535, 0); -lean_inc(x_536); -x_537 = lean_ctor_get(x_535, 1); -lean_inc(x_537); -if (lean_is_exclusive(x_535)) { - lean_ctor_release(x_535, 0); - lean_ctor_release(x_535, 1); - x_538 = x_535; +lean_object* x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; lean_object* x_527; +x_523 = lean_unsigned_to_nat(1u); +x_524 = lean_array_fget(x_290, x_523); +x_525 = lean_unsigned_to_nat(2u); +x_526 = lean_array_fget(x_290, x_525); +lean_dec(x_290); +lean_inc(x_526); +x_527 = l_Lean_Expr_int_x3f(x_526); +if (lean_obj_tag(x_527) == 0) +{ +lean_object* x_528; lean_object* x_529; lean_object* x_530; +x_528 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__22; +x_529 = l_Lean_mkApp3(x_528, x_524, x_526, x_3); +x_530 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_529, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_530; +} +else +{ +lean_object* x_531; lean_object* x_532; uint8_t x_533; +x_531 = lean_ctor_get(x_527, 0); +lean_inc(x_531); +lean_dec(x_527); +x_532 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__8; +x_533 = lean_int_dec_eq(x_531, x_532); +lean_dec(x_531); +if (x_533 == 0) +{ +lean_object* x_534; lean_object* x_535; lean_object* x_536; +x_534 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__22; +x_535 = l_Lean_mkApp3(x_534, x_524, x_526, x_3); +x_536 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_535, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_536; +} +else +{ +lean_object* x_537; +lean_dec(x_526); +x_537 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality(x_2, x_3, x_524, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_537) == 0) +{ +lean_object* x_538; lean_object* x_539; lean_object* x_540; lean_object* x_541; lean_object* x_542; +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; } else { - lean_dec_ref(x_535); - x_538 = lean_box(0); -} -x_539 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_539, 0, x_536); -lean_ctor_set(x_539, 1, x_521); -if (lean_is_scalar(x_538)) { - x_540 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_537); + x_540 = lean_box(0); +} +x_541 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_541, 0, x_538); +lean_ctor_set(x_541, 1, x_523); +if (lean_is_scalar(x_540)) { + x_542 = lean_alloc_ctor(0, 2, 0); } else { - x_540 = x_538; + x_542 = x_540; } -lean_ctor_set(x_540, 0, x_539); -lean_ctor_set(x_540, 1, x_537); -return x_540; +lean_ctor_set(x_542, 0, x_541); +lean_ctor_set(x_542, 1, x_539); +return x_542; } else { -lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; -x_541 = lean_ctor_get(x_535, 0); -lean_inc(x_541); -x_542 = lean_ctor_get(x_535, 1); -lean_inc(x_542); -if (lean_is_exclusive(x_535)) { - lean_ctor_release(x_535, 0); - lean_ctor_release(x_535, 1); - x_543 = x_535; +lean_object* x_543; lean_object* x_544; lean_object* x_545; lean_object* x_546; +x_543 = lean_ctor_get(x_537, 0); +lean_inc(x_543); +x_544 = lean_ctor_get(x_537, 1); +lean_inc(x_544); +if (lean_is_exclusive(x_537)) { + lean_ctor_release(x_537, 0); + lean_ctor_release(x_537, 1); + x_545 = x_537; } else { - lean_dec_ref(x_535); - x_543 = lean_box(0); + lean_dec_ref(x_537); + x_545 = lean_box(0); } -if (lean_is_scalar(x_543)) { - x_544 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_545)) { + x_546 = lean_alloc_ctor(1, 2, 0); } else { - x_544 = x_543; + x_546 = x_545; } -lean_ctor_set(x_544, 0, x_541); -lean_ctor_set(x_544, 1, x_542); -return x_544; +lean_ctor_set(x_546, 0, x_543); +lean_ctor_set(x_546, 1, x_544); +return x_546; } } } } else { -lean_object* x_545; lean_object* x_546; -lean_dec(x_504); -lean_dec(x_288); +lean_object* x_547; lean_object* x_548; +lean_dec(x_506); +lean_dec(x_290); +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_3); -x_545 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_545, 0, x_2); -lean_ctor_set(x_545, 1, x_500); -x_546 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_546, 0, x_545); -lean_ctor_set(x_546, 1, x_13); -return x_546; +x_547 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_547, 0, x_2); +lean_ctor_set(x_547, 1, x_502); +x_548 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_548, 0, x_547); +lean_ctor_set(x_548, 1, x_15); +return x_548; } } } else { -lean_object* x_547; lean_object* x_548; +lean_object* x_549; lean_object* x_550; +lean_dec(x_505); +lean_dec(x_504); lean_dec(x_503); -lean_dec(x_502); -lean_dec(x_501); -lean_dec(x_288); +lean_dec(x_290); +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_3); -x_547 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_547, 0, x_2); -lean_ctor_set(x_547, 1, x_500); -x_548 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_548, 0, x_547); -lean_ctor_set(x_548, 1, x_13); -return x_548; +x_549 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_549, 0, x_2); +lean_ctor_set(x_549, 1, x_502); +x_550 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_550, 0, x_549); +lean_ctor_set(x_550, 1, x_15); +return x_550; } } else { -lean_object* x_549; lean_object* x_550; -lean_dec(x_502); -lean_dec(x_501); -lean_dec(x_288); +lean_object* x_551; lean_object* x_552; +lean_dec(x_504); +lean_dec(x_503); +lean_dec(x_290); +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_3); -x_549 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_549, 0, x_2); -lean_ctor_set(x_549, 1, x_500); -x_550 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_550, 0, x_549); -lean_ctor_set(x_550, 1, x_13); -return x_550; +x_551 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_551, 0, x_2); +lean_ctor_set(x_551, 1, x_502); +x_552 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_552, 0, x_551); +lean_ctor_set(x_552, 1, x_15); +return x_552; } } case 5: { -lean_object* x_551; -x_551 = lean_ctor_get(x_501, 0); -lean_inc(x_551); -if (lean_obj_tag(x_551) == 4) -{ -lean_object* x_552; -x_552 = lean_ctor_get(x_551, 0); -lean_inc(x_552); -if (lean_obj_tag(x_552) == 1) -{ lean_object* x_553; -x_553 = lean_ctor_get(x_552, 0); +x_553 = lean_ctor_get(x_503, 0); lean_inc(x_553); -if (lean_obj_tag(x_553) == 0) +if (lean_obj_tag(x_553) == 4) { -lean_object* x_554; lean_object* x_555; lean_object* x_556; lean_object* x_557; uint8_t x_558; -x_554 = lean_ctor_get(x_501, 1); +lean_object* x_554; +x_554 = lean_ctor_get(x_553, 0); lean_inc(x_554); -lean_dec(x_501); -x_555 = lean_ctor_get(x_551, 1); +if (lean_obj_tag(x_554) == 1) +{ +lean_object* x_555; +x_555 = lean_ctor_get(x_554, 0); lean_inc(x_555); -lean_dec(x_551); -x_556 = lean_ctor_get(x_552, 1); -lean_inc(x_556); -lean_dec(x_552); -x_557 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__4; -x_558 = lean_string_dec_eq(x_556, x_557); -lean_dec(x_556); -if (x_558 == 0) +if (lean_obj_tag(x_555) == 0) { -lean_object* x_559; lean_object* x_560; -lean_dec(x_555); +lean_object* x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; uint8_t x_560; +x_556 = lean_ctor_get(x_503, 1); +lean_inc(x_556); +lean_dec(x_503); +x_557 = lean_ctor_get(x_553, 1); +lean_inc(x_557); +lean_dec(x_553); +x_558 = lean_ctor_get(x_554, 1); +lean_inc(x_558); lean_dec(x_554); -lean_dec(x_288); +x_559 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__4; +x_560 = lean_string_dec_eq(x_558, x_559); +lean_dec(x_558); +if (x_560 == 0) +{ +lean_object* x_561; lean_object* x_562; +lean_dec(x_557); +lean_dec(x_556); +lean_dec(x_290); +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_3); -x_559 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_559, 0, x_2); -lean_ctor_set(x_559, 1, x_500); -x_560 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_560, 0, x_559); -lean_ctor_set(x_560, 1, x_13); -return x_560; +x_561 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_561, 0, x_2); +lean_ctor_set(x_561, 1, x_502); +x_562 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_562, 0, x_561); +lean_ctor_set(x_562, 1, x_15); +return x_562; } else { -if (lean_obj_tag(x_555) == 0) -{ -lean_object* x_561; lean_object* x_562; lean_object* x_563; lean_object* x_564; lean_object* x_565; lean_object* x_566; lean_object* x_567; -x_561 = lean_unsigned_to_nat(1u); -x_562 = lean_array_fget(x_288, x_561); -x_563 = lean_unsigned_to_nat(2u); -x_564 = lean_array_fget(x_288, x_563); -lean_dec(x_288); -x_565 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__25; -x_566 = l_Lean_mkApp4(x_565, x_554, x_562, x_564, x_3); -x_567 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_566, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_567; -} -else +if (lean_obj_tag(x_557) == 0) { -lean_object* x_568; lean_object* x_569; -lean_dec(x_555); -lean_dec(x_554); -lean_dec(x_288); -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_3); -x_568 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_568, 0, x_2); -lean_ctor_set(x_568, 1, x_500); -x_569 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_569, 0, x_568); -lean_ctor_set(x_569, 1, x_13); +lean_object* x_563; lean_object* x_564; lean_object* x_565; lean_object* x_566; lean_object* x_567; lean_object* x_568; lean_object* x_569; +x_563 = lean_unsigned_to_nat(1u); +x_564 = lean_array_fget(x_290, x_563); +x_565 = lean_unsigned_to_nat(2u); +x_566 = lean_array_fget(x_290, x_565); +lean_dec(x_290); +x_567 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__25; +x_568 = l_Lean_mkApp4(x_567, x_556, x_564, x_566, x_3); +x_569 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_568, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); return x_569; } -} -} else { lean_object* x_570; lean_object* x_571; -lean_dec(x_553); -lean_dec(x_552); -lean_dec(x_551); -lean_dec(x_501); -lean_dec(x_288); +lean_dec(x_557); +lean_dec(x_556); +lean_dec(x_290); +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_3); x_570 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_570, 0, x_2); -lean_ctor_set(x_570, 1, x_500); +lean_ctor_set(x_570, 1, x_502); x_571 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_571, 0, x_570); -lean_ctor_set(x_571, 1, x_13); +lean_ctor_set(x_571, 1, x_15); return x_571; } } +} else { lean_object* x_572; lean_object* x_573; -lean_dec(x_552); -lean_dec(x_551); -lean_dec(x_501); -lean_dec(x_288); +lean_dec(x_555); +lean_dec(x_554); +lean_dec(x_553); +lean_dec(x_503); +lean_dec(x_290); +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_3); x_572 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_572, 0, x_2); -lean_ctor_set(x_572, 1, x_500); +lean_ctor_set(x_572, 1, x_502); x_573 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_573, 0, x_572); -lean_ctor_set(x_573, 1, x_13); +lean_ctor_set(x_573, 1, x_15); return x_573; } } else { lean_object* x_574; lean_object* x_575; -lean_dec(x_551); -lean_dec(x_501); -lean_dec(x_288); +lean_dec(x_554); +lean_dec(x_553); +lean_dec(x_503); +lean_dec(x_290); +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_3); x_574 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_574, 0, x_2); -lean_ctor_set(x_574, 1, x_500); +lean_ctor_set(x_574, 1, x_502); x_575 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_575, 0, x_574); -lean_ctor_set(x_575, 1, x_13); +lean_ctor_set(x_575, 1, x_15); return x_575; } } -default: +else { lean_object* x_576; lean_object* x_577; -lean_dec(x_501); -lean_dec(x_288); +lean_dec(x_553); +lean_dec(x_503); +lean_dec(x_290); +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_3); x_576 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_576, 0, x_2); -lean_ctor_set(x_576, 1, x_500); +lean_ctor_set(x_576, 1, x_502); x_577 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_577, 0, x_576); -lean_ctor_set(x_577, 1, x_13); +lean_ctor_set(x_577, 1, x_15); return x_577; } } +default: +{ +lean_object* x_578; lean_object* x_579; +lean_dec(x_503); +lean_dec(x_290); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_3); +x_578 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_578, 0, x_2); +lean_ctor_set(x_578, 1, x_502); +x_579 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_579, 0, x_578); +lean_ctor_set(x_579, 1, x_15); +return x_579; +} +} } } } } case 1: { -lean_object* x_578; +lean_object* x_580; lean_dec(x_4); -x_578 = lean_ctor_get(x_16, 0); -lean_inc(x_578); -if (lean_obj_tag(x_578) == 0) -{ -uint8_t x_579; -x_579 = !lean_is_exclusive(x_14); -if (x_579 == 0) -{ -lean_object* x_580; lean_object* x_581; lean_object* x_582; lean_object* x_583; lean_object* x_584; uint8_t x_585; -x_580 = lean_ctor_get(x_14, 1); -x_581 = lean_ctor_get(x_14, 0); -lean_dec(x_581); -x_582 = lean_ctor_get(x_15, 1); -lean_inc(x_582); -lean_dec(x_15); -x_583 = lean_ctor_get(x_16, 1); -lean_inc(x_583); -lean_dec(x_16); -x_584 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__10; -x_585 = lean_string_dec_eq(x_583, x_584); -if (x_585 == 0) +x_580 = lean_ctor_get(x_18, 0); +lean_inc(x_580); +if (lean_obj_tag(x_580) == 0) +{ +uint8_t x_581; +x_581 = !lean_is_exclusive(x_16); +if (x_581 == 0) { -lean_object* x_586; uint8_t x_587; -x_586 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__13; -x_587 = lean_string_dec_eq(x_583, x_586); +lean_object* x_582; lean_object* x_583; lean_object* x_584; lean_object* x_585; lean_object* x_586; uint8_t x_587; +x_582 = lean_ctor_get(x_16, 1); +x_583 = lean_ctor_get(x_16, 0); +lean_dec(x_583); +x_584 = lean_ctor_get(x_17, 1); +lean_inc(x_584); +lean_dec(x_17); +x_585 = lean_ctor_get(x_18, 1); +lean_inc(x_585); +lean_dec(x_18); +x_586 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__10; +x_587 = lean_string_dec_eq(x_585, x_586); if (x_587 == 0) { lean_object* x_588; uint8_t x_589; -x_588 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__26; -x_589 = lean_string_dec_eq(x_583, x_588); +x_588 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__13; +x_589 = lean_string_dec_eq(x_585, x_588); if (x_589 == 0) { lean_object* x_590; uint8_t x_591; -x_590 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__27; -x_591 = lean_string_dec_eq(x_583, x_590); +x_590 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__26; +x_591 = lean_string_dec_eq(x_585, x_590); if (x_591 == 0) { lean_object* x_592; uint8_t x_593; -x_592 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__8; -x_593 = lean_string_dec_eq(x_583, x_592); +x_592 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__27; +x_593 = lean_string_dec_eq(x_585, x_592); if (x_593 == 0) { lean_object* x_594; uint8_t x_595; -x_594 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__7; -x_595 = lean_string_dec_eq(x_583, x_594); -lean_dec(x_583); +x_594 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__8; +x_595 = lean_string_dec_eq(x_585, x_594); if (x_595 == 0) { -lean_object* x_596; lean_object* x_597; +lean_object* x_596; uint8_t x_597; +x_596 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__7; +x_597 = lean_string_dec_eq(x_585, x_596); +lean_dec(x_585); +if (x_597 == 0) +{ +lean_object* x_598; lean_object* x_599; +lean_dec(x_584); lean_dec(x_582); -lean_dec(x_580); +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_3); -x_596 = lean_unsigned_to_nat(0u); -lean_ctor_set(x_14, 1, x_596); -lean_ctor_set(x_14, 0, x_2); -x_597 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_597, 0, x_14); -lean_ctor_set(x_597, 1, x_13); -return x_597; +x_598 = lean_unsigned_to_nat(0u); +lean_ctor_set(x_16, 1, x_598); +lean_ctor_set(x_16, 0, x_2); +x_599 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_599, 0, x_16); +lean_ctor_set(x_599, 1, x_15); +return x_599; } else { -lean_object* x_598; uint8_t x_599; -x_598 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__8; -x_599 = lean_string_dec_eq(x_582, x_598); -lean_dec(x_582); -if (x_599 == 0) +lean_object* x_600; uint8_t x_601; +x_600 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__8; +x_601 = lean_string_dec_eq(x_584, x_600); +lean_dec(x_584); +if (x_601 == 0) { -lean_object* x_600; lean_object* x_601; -lean_dec(x_580); +lean_object* x_602; lean_object* x_603; +lean_dec(x_582); +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_3); -x_600 = lean_unsigned_to_nat(0u); -lean_ctor_set(x_14, 1, x_600); -lean_ctor_set(x_14, 0, x_2); -x_601 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_601, 0, x_14); -lean_ctor_set(x_601, 1, x_13); -return x_601; +x_602 = lean_unsigned_to_nat(0u); +lean_ctor_set(x_16, 1, x_602); +lean_ctor_set(x_16, 0, x_2); +x_603 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_603, 0, x_16); +lean_ctor_set(x_603, 1, x_15); +return x_603; } else { -lean_object* x_602; lean_object* x_603; uint8_t x_604; -x_602 = lean_array_get_size(x_580); -x_603 = lean_unsigned_to_nat(4u); -x_604 = lean_nat_dec_eq(x_602, x_603); -lean_dec(x_602); -if (x_604 == 0) +lean_object* x_604; lean_object* x_605; uint8_t x_606; +x_604 = lean_array_get_size(x_582); +x_605 = lean_unsigned_to_nat(4u); +x_606 = lean_nat_dec_eq(x_604, x_605); +lean_dec(x_604); +if (x_606 == 0) { -lean_object* x_605; lean_object* x_606; -lean_dec(x_580); +lean_object* x_607; lean_object* x_608; +lean_dec(x_582); +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_3); -x_605 = lean_unsigned_to_nat(0u); -lean_ctor_set(x_14, 1, x_605); -lean_ctor_set(x_14, 0, x_2); -x_606 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_606, 0, x_14); -lean_ctor_set(x_606, 1, x_13); -return x_606; +x_607 = lean_unsigned_to_nat(0u); +lean_ctor_set(x_16, 1, x_607); +lean_ctor_set(x_16, 0, x_2); +x_608 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_608, 0, x_16); +lean_ctor_set(x_608, 1, x_15); +return x_608; } else { -lean_object* x_607; lean_object* x_608; -x_607 = lean_unsigned_to_nat(0u); -x_608 = lean_array_fget(x_580, x_607); -if (lean_obj_tag(x_608) == 4) -{ -lean_object* x_609; -x_609 = lean_ctor_get(x_608, 0); -lean_inc(x_609); -if (lean_obj_tag(x_609) == 1) -{ -lean_object* x_610; -x_610 = lean_ctor_get(x_609, 0); -lean_inc(x_610); -if (lean_obj_tag(x_610) == 0) +lean_object* x_609; lean_object* x_610; +x_609 = lean_unsigned_to_nat(0u); +x_610 = lean_array_fget(x_582, x_609); +if (lean_obj_tag(x_610) == 4) { -lean_object* x_611; lean_object* x_612; lean_object* x_613; uint8_t x_614; -x_611 = lean_ctor_get(x_608, 1); +lean_object* x_611; +x_611 = lean_ctor_get(x_610, 0); lean_inc(x_611); -lean_dec(x_608); -x_612 = lean_ctor_get(x_609, 1); +if (lean_obj_tag(x_611) == 1) +{ +lean_object* x_612; +x_612 = lean_ctor_get(x_611, 0); lean_inc(x_612); -lean_dec(x_609); -x_613 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__1; -x_614 = lean_string_dec_eq(x_612, x_613); -if (x_614 == 0) +if (lean_obj_tag(x_612) == 0) { -lean_object* x_615; uint8_t x_616; -x_615 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__9; -x_616 = lean_string_dec_eq(x_612, x_615); -lean_dec(x_612); +lean_object* x_613; lean_object* x_614; lean_object* x_615; uint8_t x_616; +x_613 = lean_ctor_get(x_610, 1); +lean_inc(x_613); +lean_dec(x_610); +x_614 = lean_ctor_get(x_611, 1); +lean_inc(x_614); +lean_dec(x_611); +x_615 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__1; +x_616 = lean_string_dec_eq(x_614, x_615); if (x_616 == 0) { -lean_object* x_617; -lean_dec(x_611); -lean_dec(x_580); +lean_object* x_617; uint8_t x_618; +x_617 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__9; +x_618 = lean_string_dec_eq(x_614, x_617); +lean_dec(x_614); +if (x_618 == 0) +{ +lean_object* x_619; +lean_dec(x_613); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_607); -lean_ctor_set(x_14, 0, x_2); -x_617 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_617, 0, x_14); -lean_ctor_set(x_617, 1, x_13); -return x_617; +lean_ctor_set(x_16, 1, x_609); +lean_ctor_set(x_16, 0, x_2); +x_619 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_619, 0, x_16); +lean_ctor_set(x_619, 1, x_15); +return x_619; } else { -if (lean_obj_tag(x_611) == 0) +if (lean_obj_tag(x_613) == 0) { -lean_object* x_618; lean_object* x_619; lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; lean_object* x_624; -lean_free_object(x_14); -x_618 = lean_unsigned_to_nat(2u); -x_619 = lean_array_fget(x_580, x_618); -x_620 = lean_unsigned_to_nat(3u); -x_621 = lean_array_fget(x_580, x_620); -lean_dec(x_580); -x_622 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__30; -x_623 = l_Lean_mkApp3(x_622, x_619, x_621, x_3); -x_624 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_623, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_624; +lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; lean_object* x_624; lean_object* x_625; lean_object* x_626; +lean_free_object(x_16); +x_620 = lean_unsigned_to_nat(2u); +x_621 = lean_array_fget(x_582, x_620); +x_622 = lean_unsigned_to_nat(3u); +x_623 = lean_array_fget(x_582, x_622); +lean_dec(x_582); +x_624 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__30; +x_625 = l_Lean_mkApp3(x_624, x_621, x_623, x_3); +x_626 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_625, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_626; } else { -lean_object* x_625; -lean_dec(x_611); -lean_dec(x_580); +lean_object* x_627; +lean_dec(x_613); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_607); -lean_ctor_set(x_14, 0, x_2); -x_625 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_625, 0, x_14); -lean_ctor_set(x_625, 1, x_13); -return x_625; +lean_ctor_set(x_16, 1, x_609); +lean_ctor_set(x_16, 0, x_2); +x_627 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_627, 0, x_16); +lean_ctor_set(x_627, 1, x_15); +return x_627; } } } else { -lean_dec(x_612); -if (lean_obj_tag(x_611) == 0) -{ -lean_object* x_626; lean_object* x_627; lean_object* x_628; lean_object* x_629; lean_object* x_630; lean_object* x_631; lean_object* x_632; -lean_free_object(x_14); -x_626 = lean_unsigned_to_nat(2u); -x_627 = lean_array_fget(x_580, x_626); -x_628 = lean_unsigned_to_nat(3u); -x_629 = lean_array_fget(x_580, x_628); -lean_dec(x_580); -x_630 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__33; -x_631 = l_Lean_mkApp3(x_630, x_627, x_629, x_3); -x_632 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_631, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_632; +lean_dec(x_614); +if (lean_obj_tag(x_613) == 0) +{ +lean_object* x_628; lean_object* x_629; lean_object* x_630; lean_object* x_631; lean_object* x_632; lean_object* x_633; lean_object* x_634; +lean_free_object(x_16); +x_628 = lean_unsigned_to_nat(2u); +x_629 = lean_array_fget(x_582, x_628); +x_630 = lean_unsigned_to_nat(3u); +x_631 = lean_array_fget(x_582, x_630); +lean_dec(x_582); +x_632 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__33; +x_633 = l_Lean_mkApp3(x_632, x_629, x_631, x_3); +x_634 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_633, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_634; } else { -lean_object* x_633; -lean_dec(x_611); -lean_dec(x_580); +lean_object* x_635; +lean_dec(x_613); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_607); -lean_ctor_set(x_14, 0, x_2); -x_633 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_633, 0, x_14); -lean_ctor_set(x_633, 1, x_13); -return x_633; +lean_ctor_set(x_16, 1, x_609); +lean_ctor_set(x_16, 0, x_2); +x_635 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_635, 0, x_16); +lean_ctor_set(x_635, 1, x_15); +return x_635; } } } else { -lean_object* x_634; +lean_object* x_636; +lean_dec(x_612); +lean_dec(x_611); lean_dec(x_610); -lean_dec(x_609); -lean_dec(x_608); -lean_dec(x_580); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_607); -lean_ctor_set(x_14, 0, x_2); -x_634 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_634, 0, x_14); -lean_ctor_set(x_634, 1, x_13); -return x_634; +lean_ctor_set(x_16, 1, x_609); +lean_ctor_set(x_16, 0, x_2); +x_636 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_636, 0, x_16); +lean_ctor_set(x_636, 1, x_15); +return x_636; } } else { -lean_object* x_635; -lean_dec(x_609); -lean_dec(x_608); -lean_dec(x_580); +lean_object* x_637; +lean_dec(x_611); +lean_dec(x_610); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_607); -lean_ctor_set(x_14, 0, x_2); -x_635 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_635, 0, x_14); -lean_ctor_set(x_635, 1, x_13); -return x_635; +lean_ctor_set(x_16, 1, x_609); +lean_ctor_set(x_16, 0, x_2); +x_637 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_637, 0, x_16); +lean_ctor_set(x_637, 1, x_15); +return x_637; } } else { -lean_object* x_636; -lean_dec(x_608); -lean_dec(x_580); +lean_object* x_638; +lean_dec(x_610); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_607); -lean_ctor_set(x_14, 0, x_2); -x_636 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_636, 0, x_14); -lean_ctor_set(x_636, 1, x_13); -return x_636; +lean_ctor_set(x_16, 1, x_609); +lean_ctor_set(x_16, 0, x_2); +x_638 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_638, 0, x_16); +lean_ctor_set(x_638, 1, x_15); +return x_638; } } } @@ -22491,82 +22847,84 @@ return x_636; } else { -lean_object* x_637; uint8_t x_638; -lean_dec(x_583); -lean_dec(x_580); -x_637 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__16; -x_638 = lean_string_dec_eq(x_582, x_637); +lean_object* x_639; uint8_t x_640; +lean_dec(x_585); lean_dec(x_582); -if (x_638 == 0) +x_639 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__16; +x_640 = lean_string_dec_eq(x_584, x_639); +lean_dec(x_584); +if (x_640 == 0) { -lean_object* x_639; lean_object* x_640; +lean_object* x_641; lean_object* x_642; +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_3); -x_639 = lean_unsigned_to_nat(0u); -lean_ctor_set(x_14, 1, x_639); -lean_ctor_set(x_14, 0, x_2); -x_640 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_640, 0, x_14); -lean_ctor_set(x_640, 1, x_13); -return x_640; +x_641 = lean_unsigned_to_nat(0u); +lean_ctor_set(x_16, 1, x_641); +lean_ctor_set(x_16, 0, x_2); +x_642 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_642, 0, x_16); +lean_ctor_set(x_642, 1, x_15); +return x_642; } else { -lean_object* x_641; lean_object* x_642; lean_object* x_643; lean_object* x_644; -lean_free_object(x_14); -x_641 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__5___closed__6; -x_642 = lean_array_push(x_641, x_3); -x_643 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__35; +lean_object* x_643; lean_object* x_644; lean_object* x_645; lean_object* x_646; +lean_free_object(x_16); +x_643 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__5___closed__6; +x_644 = lean_array_push(x_643, x_3); +x_645 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__35; +lean_inc(x_14); +lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -x_644 = l_Lean_Meta_mkAppM(x_643, x_642, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_644) == 0) +x_646 = l_Lean_Meta_mkAppM(x_645, x_644, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_646) == 0) { -lean_object* x_645; lean_object* x_646; lean_object* x_647; -x_645 = lean_ctor_get(x_644, 0); -lean_inc(x_645); -x_646 = lean_ctor_get(x_644, 1); -lean_inc(x_646); -lean_dec(x_644); -x_647 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_645, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_646); -return x_647; +lean_object* x_647; lean_object* x_648; lean_object* x_649; +x_647 = lean_ctor_get(x_646, 0); +lean_inc(x_647); +x_648 = lean_ctor_get(x_646, 1); +lean_inc(x_648); +lean_dec(x_646); +x_649 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_647, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_648); +return x_649; } else { -uint8_t x_648; +uint8_t x_650; +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_2); -x_648 = !lean_is_exclusive(x_644); -if (x_648 == 0) +x_650 = !lean_is_exclusive(x_646); +if (x_650 == 0) { -return x_644; +return x_646; } else { -lean_object* x_649; lean_object* x_650; lean_object* x_651; -x_649 = lean_ctor_get(x_644, 0); -x_650 = lean_ctor_get(x_644, 1); -lean_inc(x_650); -lean_inc(x_649); -lean_dec(x_644); -x_651 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_651, 0, x_649); -lean_ctor_set(x_651, 1, x_650); -return x_651; +lean_object* x_651; lean_object* x_652; lean_object* x_653; +x_651 = lean_ctor_get(x_646, 0); +x_652 = lean_ctor_get(x_646, 1); +lean_inc(x_652); +lean_inc(x_651); +lean_dec(x_646); +x_653 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_653, 0, x_651); +lean_ctor_set(x_653, 1, x_652); +return x_653; } } } @@ -22574,411 +22932,424 @@ return x_651; } else { -lean_object* x_652; uint8_t x_653; -lean_dec(x_583); -x_652 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__36; -x_653 = lean_string_dec_eq(x_582, x_652); -lean_dec(x_582); -if (x_653 == 0) +lean_object* x_654; uint8_t x_655; +lean_dec(x_585); +x_654 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__36; +x_655 = lean_string_dec_eq(x_584, x_654); +lean_dec(x_584); +if (x_655 == 0) { -lean_object* x_654; lean_object* x_655; -lean_dec(x_580); +lean_object* x_656; lean_object* x_657; +lean_dec(x_582); +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_3); -x_654 = lean_unsigned_to_nat(0u); -lean_ctor_set(x_14, 1, x_654); -lean_ctor_set(x_14, 0, x_2); -x_655 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_655, 0, x_14); -lean_ctor_set(x_655, 1, x_13); -return x_655; +x_656 = lean_unsigned_to_nat(0u); +lean_ctor_set(x_16, 1, x_656); +lean_ctor_set(x_16, 0, x_2); +x_657 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_657, 0, x_16); +lean_ctor_set(x_657, 1, x_15); +return x_657; } else { -lean_object* x_656; lean_object* x_657; uint8_t x_658; -x_656 = lean_array_get_size(x_580); -x_657 = lean_unsigned_to_nat(4u); -x_658 = lean_nat_dec_eq(x_656, x_657); -lean_dec(x_656); -if (x_658 == 0) +lean_object* x_658; lean_object* x_659; uint8_t x_660; +x_658 = lean_array_get_size(x_582); +x_659 = lean_unsigned_to_nat(4u); +x_660 = lean_nat_dec_eq(x_658, x_659); +lean_dec(x_658); +if (x_660 == 0) { -lean_object* x_659; lean_object* x_660; -lean_dec(x_580); +lean_object* x_661; lean_object* x_662; +lean_dec(x_582); +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_3); -x_659 = lean_unsigned_to_nat(0u); -lean_ctor_set(x_14, 1, x_659); -lean_ctor_set(x_14, 0, x_2); -x_660 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_660, 0, x_14); -lean_ctor_set(x_660, 1, x_13); -return x_660; +x_661 = lean_unsigned_to_nat(0u); +lean_ctor_set(x_16, 1, x_661); +lean_ctor_set(x_16, 0, x_2); +x_662 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_662, 0, x_16); +lean_ctor_set(x_662, 1, x_15); +return x_662; } else { -lean_object* x_661; lean_object* x_662; -x_661 = lean_unsigned_to_nat(0u); -x_662 = lean_array_fget(x_580, x_661); -switch (lean_obj_tag(x_662)) { +lean_object* x_663; lean_object* x_664; +x_663 = lean_unsigned_to_nat(0u); +x_664 = lean_array_fget(x_582, x_663); +switch (lean_obj_tag(x_664)) { case 4: { -lean_object* x_663; -x_663 = lean_ctor_get(x_662, 0); -lean_inc(x_663); -if (lean_obj_tag(x_663) == 1) -{ -lean_object* x_664; -x_664 = lean_ctor_get(x_663, 0); -lean_inc(x_664); -if (lean_obj_tag(x_664) == 0) -{ -lean_object* x_665; lean_object* x_666; lean_object* x_667; uint8_t x_668; -x_665 = lean_ctor_get(x_662, 1); +lean_object* x_665; +x_665 = lean_ctor_get(x_664, 0); lean_inc(x_665); -lean_dec(x_662); -x_666 = lean_ctor_get(x_663, 1); +if (lean_obj_tag(x_665) == 1) +{ +lean_object* x_666; +x_666 = lean_ctor_get(x_665, 0); lean_inc(x_666); -lean_dec(x_663); -x_667 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__9; -x_668 = lean_string_dec_eq(x_666, x_667); -if (x_668 == 0) -{ -lean_object* x_669; uint8_t x_670; -x_669 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__1; -x_670 = lean_string_dec_eq(x_666, x_669); -lean_dec(x_666); -if (x_670 == 0) +if (lean_obj_tag(x_666) == 0) { -lean_object* x_671; +lean_object* x_667; lean_object* x_668; lean_object* x_669; uint8_t x_670; +x_667 = lean_ctor_get(x_664, 1); +lean_inc(x_667); +lean_dec(x_664); +x_668 = lean_ctor_get(x_665, 1); +lean_inc(x_668); lean_dec(x_665); -lean_dec(x_580); +x_669 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__9; +x_670 = lean_string_dec_eq(x_668, x_669); +if (x_670 == 0) +{ +lean_object* x_671; uint8_t x_672; +x_671 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__1; +x_672 = lean_string_dec_eq(x_668, x_671); +lean_dec(x_668); +if (x_672 == 0) +{ +lean_object* x_673; +lean_dec(x_667); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_661); -lean_ctor_set(x_14, 0, x_2); -x_671 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_671, 0, x_14); -lean_ctor_set(x_671, 1, x_13); -return x_671; +lean_ctor_set(x_16, 1, x_663); +lean_ctor_set(x_16, 0, x_2); +x_673 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_673, 0, x_16); +lean_ctor_set(x_673, 1, x_15); +return x_673; } else { -if (lean_obj_tag(x_665) == 0) +if (lean_obj_tag(x_667) == 0) { -lean_object* x_672; lean_object* x_673; lean_object* x_674; lean_object* x_675; lean_object* x_676; lean_object* x_677; lean_object* x_678; -lean_free_object(x_14); -x_672 = lean_unsigned_to_nat(2u); -x_673 = lean_array_fget(x_580, x_672); -x_674 = lean_unsigned_to_nat(3u); -x_675 = lean_array_fget(x_580, x_674); -lean_dec(x_580); -x_676 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__39; -x_677 = l_Lean_mkApp3(x_676, x_673, x_675, x_3); -x_678 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_677, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_678; +lean_object* x_674; lean_object* x_675; lean_object* x_676; lean_object* x_677; lean_object* x_678; lean_object* x_679; lean_object* x_680; +lean_free_object(x_16); +x_674 = lean_unsigned_to_nat(2u); +x_675 = lean_array_fget(x_582, x_674); +x_676 = lean_unsigned_to_nat(3u); +x_677 = lean_array_fget(x_582, x_676); +lean_dec(x_582); +x_678 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__39; +x_679 = l_Lean_mkApp3(x_678, x_675, x_677, x_3); +x_680 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_679, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_680; } else { -lean_object* x_679; -lean_dec(x_665); -lean_dec(x_580); +lean_object* x_681; +lean_dec(x_667); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_661); -lean_ctor_set(x_14, 0, x_2); -x_679 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_679, 0, x_14); -lean_ctor_set(x_679, 1, x_13); -return x_679; +lean_ctor_set(x_16, 1, x_663); +lean_ctor_set(x_16, 0, x_2); +x_681 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_681, 0, x_16); +lean_ctor_set(x_681, 1, x_15); +return x_681; } } } else { -lean_dec(x_666); -if (lean_obj_tag(x_665) == 0) -{ -lean_object* x_680; lean_object* x_681; lean_object* x_682; lean_object* x_683; lean_object* x_684; lean_object* x_685; lean_object* x_686; -lean_free_object(x_14); -x_680 = lean_unsigned_to_nat(2u); -x_681 = lean_array_fget(x_580, x_680); -x_682 = lean_unsigned_to_nat(3u); -x_683 = lean_array_fget(x_580, x_682); -lean_dec(x_580); -x_684 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__41; -x_685 = l_Lean_mkApp3(x_684, x_681, x_683, x_3); -x_686 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_685, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_686; +lean_dec(x_668); +if (lean_obj_tag(x_667) == 0) +{ +lean_object* x_682; lean_object* x_683; lean_object* x_684; lean_object* x_685; lean_object* x_686; lean_object* x_687; lean_object* x_688; +lean_free_object(x_16); +x_682 = lean_unsigned_to_nat(2u); +x_683 = lean_array_fget(x_582, x_682); +x_684 = lean_unsigned_to_nat(3u); +x_685 = lean_array_fget(x_582, x_684); +lean_dec(x_582); +x_686 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__41; +x_687 = l_Lean_mkApp3(x_686, x_683, x_685, x_3); +x_688 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_687, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_688; } else { -lean_object* x_687; -lean_dec(x_665); -lean_dec(x_580); +lean_object* x_689; +lean_dec(x_667); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_661); -lean_ctor_set(x_14, 0, x_2); -x_687 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_687, 0, x_14); -lean_ctor_set(x_687, 1, x_13); -return x_687; +lean_ctor_set(x_16, 1, x_663); +lean_ctor_set(x_16, 0, x_2); +x_689 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_689, 0, x_16); +lean_ctor_set(x_689, 1, x_15); +return x_689; } } } else { -lean_object* x_688; +lean_object* x_690; +lean_dec(x_666); +lean_dec(x_665); lean_dec(x_664); -lean_dec(x_663); -lean_dec(x_662); -lean_dec(x_580); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_661); -lean_ctor_set(x_14, 0, x_2); -x_688 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_688, 0, x_14); -lean_ctor_set(x_688, 1, x_13); -return x_688; +lean_ctor_set(x_16, 1, x_663); +lean_ctor_set(x_16, 0, x_2); +x_690 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_690, 0, x_16); +lean_ctor_set(x_690, 1, x_15); +return x_690; } } else { -lean_object* x_689; -lean_dec(x_663); -lean_dec(x_662); -lean_dec(x_580); +lean_object* x_691; +lean_dec(x_665); +lean_dec(x_664); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_661); -lean_ctor_set(x_14, 0, x_2); -x_689 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_689, 0, x_14); -lean_ctor_set(x_689, 1, x_13); -return x_689; +lean_ctor_set(x_16, 1, x_663); +lean_ctor_set(x_16, 0, x_2); +x_691 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_691, 0, x_16); +lean_ctor_set(x_691, 1, x_15); +return x_691; } } case 5: { -lean_object* x_690; -x_690 = lean_ctor_get(x_662, 0); -lean_inc(x_690); -if (lean_obj_tag(x_690) == 4) -{ -lean_object* x_691; -x_691 = lean_ctor_get(x_690, 0); -lean_inc(x_691); -if (lean_obj_tag(x_691) == 1) -{ lean_object* x_692; -x_692 = lean_ctor_get(x_691, 0); +x_692 = lean_ctor_get(x_664, 0); lean_inc(x_692); -if (lean_obj_tag(x_692) == 0) +if (lean_obj_tag(x_692) == 4) { -lean_object* x_693; lean_object* x_694; lean_object* x_695; lean_object* x_696; uint8_t x_697; -x_693 = lean_ctor_get(x_662, 1); +lean_object* x_693; +x_693 = lean_ctor_get(x_692, 0); lean_inc(x_693); -lean_dec(x_662); -x_694 = lean_ctor_get(x_690, 1); +if (lean_obj_tag(x_693) == 1) +{ +lean_object* x_694; +x_694 = lean_ctor_get(x_693, 0); lean_inc(x_694); -lean_dec(x_690); -x_695 = lean_ctor_get(x_691, 1); -lean_inc(x_695); -lean_dec(x_691); -x_696 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__4; -x_697 = lean_string_dec_eq(x_695, x_696); -lean_dec(x_695); -if (x_697 == 0) +if (lean_obj_tag(x_694) == 0) { -lean_object* x_698; -lean_dec(x_694); +lean_object* x_695; lean_object* x_696; lean_object* x_697; lean_object* x_698; uint8_t x_699; +x_695 = lean_ctor_get(x_664, 1); +lean_inc(x_695); +lean_dec(x_664); +x_696 = lean_ctor_get(x_692, 1); +lean_inc(x_696); +lean_dec(x_692); +x_697 = lean_ctor_get(x_693, 1); +lean_inc(x_697); lean_dec(x_693); -lean_dec(x_580); +x_698 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__4; +x_699 = lean_string_dec_eq(x_697, x_698); +lean_dec(x_697); +if (x_699 == 0) +{ +lean_object* x_700; +lean_dec(x_696); +lean_dec(x_695); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_661); -lean_ctor_set(x_14, 0, x_2); -x_698 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_698, 0, x_14); -lean_ctor_set(x_698, 1, x_13); -return x_698; +lean_ctor_set(x_16, 1, x_663); +lean_ctor_set(x_16, 0, x_2); +x_700 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_700, 0, x_16); +lean_ctor_set(x_700, 1, x_15); +return x_700; } else { -if (lean_obj_tag(x_694) == 0) +if (lean_obj_tag(x_696) == 0) { -lean_object* x_699; lean_object* x_700; lean_object* x_701; lean_object* x_702; lean_object* x_703; lean_object* x_704; lean_object* x_705; -lean_free_object(x_14); -x_699 = lean_unsigned_to_nat(2u); -x_700 = lean_array_fget(x_580, x_699); -x_701 = lean_unsigned_to_nat(3u); -x_702 = lean_array_fget(x_580, x_701); -lean_dec(x_580); -x_703 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__44; -x_704 = l_Lean_mkApp4(x_703, x_693, x_700, x_702, x_3); -x_705 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_704, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_705; +lean_object* x_701; lean_object* x_702; lean_object* x_703; lean_object* x_704; lean_object* x_705; lean_object* x_706; lean_object* x_707; +lean_free_object(x_16); +x_701 = lean_unsigned_to_nat(2u); +x_702 = lean_array_fget(x_582, x_701); +x_703 = lean_unsigned_to_nat(3u); +x_704 = lean_array_fget(x_582, x_703); +lean_dec(x_582); +x_705 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__44; +x_706 = l_Lean_mkApp4(x_705, x_695, x_702, x_704, x_3); +x_707 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_706, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_707; } else { -lean_object* x_706; -lean_dec(x_694); -lean_dec(x_693); -lean_dec(x_580); +lean_object* x_708; +lean_dec(x_696); +lean_dec(x_695); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_661); -lean_ctor_set(x_14, 0, x_2); -x_706 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_706, 0, x_14); -lean_ctor_set(x_706, 1, x_13); -return x_706; +lean_ctor_set(x_16, 1, x_663); +lean_ctor_set(x_16, 0, x_2); +x_708 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_708, 0, x_16); +lean_ctor_set(x_708, 1, x_15); +return x_708; } } } else { -lean_object* x_707; +lean_object* x_709; +lean_dec(x_694); +lean_dec(x_693); lean_dec(x_692); -lean_dec(x_691); -lean_dec(x_690); -lean_dec(x_662); -lean_dec(x_580); +lean_dec(x_664); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_661); -lean_ctor_set(x_14, 0, x_2); -x_707 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_707, 0, x_14); -lean_ctor_set(x_707, 1, x_13); -return x_707; +lean_ctor_set(x_16, 1, x_663); +lean_ctor_set(x_16, 0, x_2); +x_709 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_709, 0, x_16); +lean_ctor_set(x_709, 1, x_15); +return x_709; } } else { -lean_object* x_708; -lean_dec(x_691); -lean_dec(x_690); -lean_dec(x_662); -lean_dec(x_580); +lean_object* x_710; +lean_dec(x_693); +lean_dec(x_692); +lean_dec(x_664); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_661); -lean_ctor_set(x_14, 0, x_2); -x_708 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_708, 0, x_14); -lean_ctor_set(x_708, 1, x_13); -return x_708; +lean_ctor_set(x_16, 1, x_663); +lean_ctor_set(x_16, 0, x_2); +x_710 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_710, 0, x_16); +lean_ctor_set(x_710, 1, x_15); +return x_710; } } else { -lean_object* x_709; -lean_dec(x_690); -lean_dec(x_662); -lean_dec(x_580); +lean_object* x_711; +lean_dec(x_692); +lean_dec(x_664); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_661); -lean_ctor_set(x_14, 0, x_2); -x_709 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_709, 0, x_14); -lean_ctor_set(x_709, 1, x_13); -return x_709; +lean_ctor_set(x_16, 1, x_663); +lean_ctor_set(x_16, 0, x_2); +x_711 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_711, 0, x_16); +lean_ctor_set(x_711, 1, x_15); +return x_711; } } default: { -lean_object* x_710; -lean_dec(x_662); -lean_dec(x_580); +lean_object* x_712; +lean_dec(x_664); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_661); -lean_ctor_set(x_14, 0, x_2); -x_710 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_710, 0, x_14); -lean_ctor_set(x_710, 1, x_13); -return x_710; +lean_ctor_set(x_16, 1, x_663); +lean_ctor_set(x_16, 0, x_2); +x_712 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_712, 0, x_16); +lean_ctor_set(x_712, 1, x_15); +return x_712; } } } @@ -22987,411 +23358,424 @@ return x_710; } else { -lean_object* x_711; uint8_t x_712; -lean_dec(x_583); -x_711 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__45; -x_712 = lean_string_dec_eq(x_582, x_711); -lean_dec(x_582); -if (x_712 == 0) +lean_object* x_713; uint8_t x_714; +lean_dec(x_585); +x_713 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__45; +x_714 = lean_string_dec_eq(x_584, x_713); +lean_dec(x_584); +if (x_714 == 0) { -lean_object* x_713; lean_object* x_714; -lean_dec(x_580); +lean_object* x_715; lean_object* x_716; +lean_dec(x_582); +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_3); -x_713 = lean_unsigned_to_nat(0u); -lean_ctor_set(x_14, 1, x_713); -lean_ctor_set(x_14, 0, x_2); -x_714 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_714, 0, x_14); -lean_ctor_set(x_714, 1, x_13); -return x_714; +x_715 = lean_unsigned_to_nat(0u); +lean_ctor_set(x_16, 1, x_715); +lean_ctor_set(x_16, 0, x_2); +x_716 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_716, 0, x_16); +lean_ctor_set(x_716, 1, x_15); +return x_716; } else { -lean_object* x_715; lean_object* x_716; uint8_t x_717; -x_715 = lean_array_get_size(x_580); -x_716 = lean_unsigned_to_nat(4u); -x_717 = lean_nat_dec_eq(x_715, x_716); -lean_dec(x_715); -if (x_717 == 0) +lean_object* x_717; lean_object* x_718; uint8_t x_719; +x_717 = lean_array_get_size(x_582); +x_718 = lean_unsigned_to_nat(4u); +x_719 = lean_nat_dec_eq(x_717, x_718); +lean_dec(x_717); +if (x_719 == 0) { -lean_object* x_718; lean_object* x_719; -lean_dec(x_580); +lean_object* x_720; lean_object* x_721; +lean_dec(x_582); +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_3); -x_718 = lean_unsigned_to_nat(0u); -lean_ctor_set(x_14, 1, x_718); -lean_ctor_set(x_14, 0, x_2); -x_719 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_719, 0, x_14); -lean_ctor_set(x_719, 1, x_13); -return x_719; +x_720 = lean_unsigned_to_nat(0u); +lean_ctor_set(x_16, 1, x_720); +lean_ctor_set(x_16, 0, x_2); +x_721 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_721, 0, x_16); +lean_ctor_set(x_721, 1, x_15); +return x_721; } else { -lean_object* x_720; lean_object* x_721; -x_720 = lean_unsigned_to_nat(0u); -x_721 = lean_array_fget(x_580, x_720); -switch (lean_obj_tag(x_721)) { +lean_object* x_722; lean_object* x_723; +x_722 = lean_unsigned_to_nat(0u); +x_723 = lean_array_fget(x_582, x_722); +switch (lean_obj_tag(x_723)) { case 4: { -lean_object* x_722; -x_722 = lean_ctor_get(x_721, 0); -lean_inc(x_722); -if (lean_obj_tag(x_722) == 1) -{ -lean_object* x_723; -x_723 = lean_ctor_get(x_722, 0); -lean_inc(x_723); -if (lean_obj_tag(x_723) == 0) -{ -lean_object* x_724; lean_object* x_725; lean_object* x_726; uint8_t x_727; -x_724 = lean_ctor_get(x_721, 1); +lean_object* x_724; +x_724 = lean_ctor_get(x_723, 0); lean_inc(x_724); -lean_dec(x_721); -x_725 = lean_ctor_get(x_722, 1); +if (lean_obj_tag(x_724) == 1) +{ +lean_object* x_725; +x_725 = lean_ctor_get(x_724, 0); lean_inc(x_725); -lean_dec(x_722); -x_726 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__9; -x_727 = lean_string_dec_eq(x_725, x_726); -if (x_727 == 0) -{ -lean_object* x_728; uint8_t x_729; -x_728 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__1; -x_729 = lean_string_dec_eq(x_725, x_728); -lean_dec(x_725); -if (x_729 == 0) +if (lean_obj_tag(x_725) == 0) { -lean_object* x_730; +lean_object* x_726; lean_object* x_727; lean_object* x_728; uint8_t x_729; +x_726 = lean_ctor_get(x_723, 1); +lean_inc(x_726); +lean_dec(x_723); +x_727 = lean_ctor_get(x_724, 1); +lean_inc(x_727); lean_dec(x_724); -lean_dec(x_580); +x_728 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__9; +x_729 = lean_string_dec_eq(x_727, x_728); +if (x_729 == 0) +{ +lean_object* x_730; uint8_t x_731; +x_730 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__1; +x_731 = lean_string_dec_eq(x_727, x_730); +lean_dec(x_727); +if (x_731 == 0) +{ +lean_object* x_732; +lean_dec(x_726); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_720); -lean_ctor_set(x_14, 0, x_2); -x_730 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_730, 0, x_14); -lean_ctor_set(x_730, 1, x_13); -return x_730; +lean_ctor_set(x_16, 1, x_722); +lean_ctor_set(x_16, 0, x_2); +x_732 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_732, 0, x_16); +lean_ctor_set(x_732, 1, x_15); +return x_732; } else { -if (lean_obj_tag(x_724) == 0) +if (lean_obj_tag(x_726) == 0) { -lean_object* x_731; lean_object* x_732; lean_object* x_733; lean_object* x_734; lean_object* x_735; lean_object* x_736; lean_object* x_737; -lean_free_object(x_14); -x_731 = lean_unsigned_to_nat(2u); -x_732 = lean_array_fget(x_580, x_731); -x_733 = lean_unsigned_to_nat(3u); -x_734 = lean_array_fget(x_580, x_733); -lean_dec(x_580); -x_735 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__48; -x_736 = l_Lean_mkApp3(x_735, x_732, x_734, x_3); -x_737 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_736, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_737; +lean_object* x_733; lean_object* x_734; lean_object* x_735; lean_object* x_736; lean_object* x_737; lean_object* x_738; lean_object* x_739; +lean_free_object(x_16); +x_733 = lean_unsigned_to_nat(2u); +x_734 = lean_array_fget(x_582, x_733); +x_735 = lean_unsigned_to_nat(3u); +x_736 = lean_array_fget(x_582, x_735); +lean_dec(x_582); +x_737 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__48; +x_738 = l_Lean_mkApp3(x_737, x_734, x_736, x_3); +x_739 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_738, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_739; } else { -lean_object* x_738; -lean_dec(x_724); -lean_dec(x_580); +lean_object* x_740; +lean_dec(x_726); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_720); -lean_ctor_set(x_14, 0, x_2); -x_738 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_738, 0, x_14); -lean_ctor_set(x_738, 1, x_13); -return x_738; +lean_ctor_set(x_16, 1, x_722); +lean_ctor_set(x_16, 0, x_2); +x_740 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_740, 0, x_16); +lean_ctor_set(x_740, 1, x_15); +return x_740; } } } else { -lean_dec(x_725); -if (lean_obj_tag(x_724) == 0) -{ -lean_object* x_739; lean_object* x_740; lean_object* x_741; lean_object* x_742; lean_object* x_743; lean_object* x_744; lean_object* x_745; -lean_free_object(x_14); -x_739 = lean_unsigned_to_nat(2u); -x_740 = lean_array_fget(x_580, x_739); -x_741 = lean_unsigned_to_nat(3u); -x_742 = lean_array_fget(x_580, x_741); -lean_dec(x_580); -x_743 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__50; -x_744 = l_Lean_mkApp3(x_743, x_740, x_742, x_3); -x_745 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_744, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_745; +lean_dec(x_727); +if (lean_obj_tag(x_726) == 0) +{ +lean_object* x_741; lean_object* x_742; lean_object* x_743; lean_object* x_744; lean_object* x_745; lean_object* x_746; lean_object* x_747; +lean_free_object(x_16); +x_741 = lean_unsigned_to_nat(2u); +x_742 = lean_array_fget(x_582, x_741); +x_743 = lean_unsigned_to_nat(3u); +x_744 = lean_array_fget(x_582, x_743); +lean_dec(x_582); +x_745 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__50; +x_746 = l_Lean_mkApp3(x_745, x_742, x_744, x_3); +x_747 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_746, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_747; } else { -lean_object* x_746; -lean_dec(x_724); -lean_dec(x_580); +lean_object* x_748; +lean_dec(x_726); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_720); -lean_ctor_set(x_14, 0, x_2); -x_746 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_746, 0, x_14); -lean_ctor_set(x_746, 1, x_13); -return x_746; +lean_ctor_set(x_16, 1, x_722); +lean_ctor_set(x_16, 0, x_2); +x_748 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_748, 0, x_16); +lean_ctor_set(x_748, 1, x_15); +return x_748; } } } else { -lean_object* x_747; +lean_object* x_749; +lean_dec(x_725); +lean_dec(x_724); lean_dec(x_723); -lean_dec(x_722); -lean_dec(x_721); -lean_dec(x_580); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_720); -lean_ctor_set(x_14, 0, x_2); -x_747 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_747, 0, x_14); -lean_ctor_set(x_747, 1, x_13); -return x_747; +lean_ctor_set(x_16, 1, x_722); +lean_ctor_set(x_16, 0, x_2); +x_749 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_749, 0, x_16); +lean_ctor_set(x_749, 1, x_15); +return x_749; } } else { -lean_object* x_748; -lean_dec(x_722); -lean_dec(x_721); -lean_dec(x_580); +lean_object* x_750; +lean_dec(x_724); +lean_dec(x_723); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_720); -lean_ctor_set(x_14, 0, x_2); -x_748 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_748, 0, x_14); -lean_ctor_set(x_748, 1, x_13); -return x_748; +lean_ctor_set(x_16, 1, x_722); +lean_ctor_set(x_16, 0, x_2); +x_750 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_750, 0, x_16); +lean_ctor_set(x_750, 1, x_15); +return x_750; } } case 5: { -lean_object* x_749; -x_749 = lean_ctor_get(x_721, 0); -lean_inc(x_749); -if (lean_obj_tag(x_749) == 4) -{ -lean_object* x_750; -x_750 = lean_ctor_get(x_749, 0); -lean_inc(x_750); -if (lean_obj_tag(x_750) == 1) -{ lean_object* x_751; -x_751 = lean_ctor_get(x_750, 0); +x_751 = lean_ctor_get(x_723, 0); lean_inc(x_751); -if (lean_obj_tag(x_751) == 0) +if (lean_obj_tag(x_751) == 4) { -lean_object* x_752; lean_object* x_753; lean_object* x_754; lean_object* x_755; uint8_t x_756; -x_752 = lean_ctor_get(x_721, 1); +lean_object* x_752; +x_752 = lean_ctor_get(x_751, 0); lean_inc(x_752); -lean_dec(x_721); -x_753 = lean_ctor_get(x_749, 1); +if (lean_obj_tag(x_752) == 1) +{ +lean_object* x_753; +x_753 = lean_ctor_get(x_752, 0); lean_inc(x_753); -lean_dec(x_749); -x_754 = lean_ctor_get(x_750, 1); -lean_inc(x_754); -lean_dec(x_750); -x_755 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__4; -x_756 = lean_string_dec_eq(x_754, x_755); -lean_dec(x_754); -if (x_756 == 0) +if (lean_obj_tag(x_753) == 0) { -lean_object* x_757; -lean_dec(x_753); +lean_object* x_754; lean_object* x_755; lean_object* x_756; lean_object* x_757; uint8_t x_758; +x_754 = lean_ctor_get(x_723, 1); +lean_inc(x_754); +lean_dec(x_723); +x_755 = lean_ctor_get(x_751, 1); +lean_inc(x_755); +lean_dec(x_751); +x_756 = lean_ctor_get(x_752, 1); +lean_inc(x_756); lean_dec(x_752); -lean_dec(x_580); +x_757 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__4; +x_758 = lean_string_dec_eq(x_756, x_757); +lean_dec(x_756); +if (x_758 == 0) +{ +lean_object* x_759; +lean_dec(x_755); +lean_dec(x_754); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_720); -lean_ctor_set(x_14, 0, x_2); -x_757 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_757, 0, x_14); -lean_ctor_set(x_757, 1, x_13); -return x_757; +lean_ctor_set(x_16, 1, x_722); +lean_ctor_set(x_16, 0, x_2); +x_759 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_759, 0, x_16); +lean_ctor_set(x_759, 1, x_15); +return x_759; } else { -if (lean_obj_tag(x_753) == 0) +if (lean_obj_tag(x_755) == 0) { -lean_object* x_758; lean_object* x_759; lean_object* x_760; lean_object* x_761; lean_object* x_762; lean_object* x_763; lean_object* x_764; -lean_free_object(x_14); -x_758 = lean_unsigned_to_nat(2u); -x_759 = lean_array_fget(x_580, x_758); -x_760 = lean_unsigned_to_nat(3u); -x_761 = lean_array_fget(x_580, x_760); -lean_dec(x_580); -x_762 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__53; -x_763 = l_Lean_mkApp4(x_762, x_752, x_759, x_761, x_3); -x_764 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_763, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_764; +lean_object* x_760; lean_object* x_761; lean_object* x_762; lean_object* x_763; lean_object* x_764; lean_object* x_765; lean_object* x_766; +lean_free_object(x_16); +x_760 = lean_unsigned_to_nat(2u); +x_761 = lean_array_fget(x_582, x_760); +x_762 = lean_unsigned_to_nat(3u); +x_763 = lean_array_fget(x_582, x_762); +lean_dec(x_582); +x_764 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__53; +x_765 = l_Lean_mkApp4(x_764, x_754, x_761, x_763, x_3); +x_766 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_765, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_766; } else { -lean_object* x_765; -lean_dec(x_753); -lean_dec(x_752); -lean_dec(x_580); +lean_object* x_767; +lean_dec(x_755); +lean_dec(x_754); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_720); -lean_ctor_set(x_14, 0, x_2); -x_765 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_765, 0, x_14); -lean_ctor_set(x_765, 1, x_13); -return x_765; +lean_ctor_set(x_16, 1, x_722); +lean_ctor_set(x_16, 0, x_2); +x_767 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_767, 0, x_16); +lean_ctor_set(x_767, 1, x_15); +return x_767; } } } else { -lean_object* x_766; +lean_object* x_768; +lean_dec(x_753); +lean_dec(x_752); lean_dec(x_751); -lean_dec(x_750); -lean_dec(x_749); -lean_dec(x_721); -lean_dec(x_580); +lean_dec(x_723); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_720); -lean_ctor_set(x_14, 0, x_2); -x_766 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_766, 0, x_14); -lean_ctor_set(x_766, 1, x_13); -return x_766; +lean_ctor_set(x_16, 1, x_722); +lean_ctor_set(x_16, 0, x_2); +x_768 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_768, 0, x_16); +lean_ctor_set(x_768, 1, x_15); +return x_768; } } else { -lean_object* x_767; -lean_dec(x_750); -lean_dec(x_749); -lean_dec(x_721); -lean_dec(x_580); +lean_object* x_769; +lean_dec(x_752); +lean_dec(x_751); +lean_dec(x_723); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_720); -lean_ctor_set(x_14, 0, x_2); -x_767 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_767, 0, x_14); -lean_ctor_set(x_767, 1, x_13); -return x_767; +lean_ctor_set(x_16, 1, x_722); +lean_ctor_set(x_16, 0, x_2); +x_769 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_769, 0, x_16); +lean_ctor_set(x_769, 1, x_15); +return x_769; } } else { -lean_object* x_768; -lean_dec(x_749); -lean_dec(x_721); -lean_dec(x_580); +lean_object* x_770; +lean_dec(x_751); +lean_dec(x_723); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_720); -lean_ctor_set(x_14, 0, x_2); -x_768 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_768, 0, x_14); -lean_ctor_set(x_768, 1, x_13); -return x_768; +lean_ctor_set(x_16, 1, x_722); +lean_ctor_set(x_16, 0, x_2); +x_770 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_770, 0, x_16); +lean_ctor_set(x_770, 1, x_15); +return x_770; } } default: { -lean_object* x_769; -lean_dec(x_721); -lean_dec(x_580); +lean_object* x_771; +lean_dec(x_723); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_720); -lean_ctor_set(x_14, 0, x_2); -x_769 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_769, 0, x_14); -lean_ctor_set(x_769, 1, x_13); -return x_769; +lean_ctor_set(x_16, 1, x_722); +lean_ctor_set(x_16, 0, x_2); +x_771 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_771, 0, x_16); +lean_ctor_set(x_771, 1, x_15); +return x_771; } } } @@ -23400,411 +23784,424 @@ return x_769; } else { -lean_object* x_770; uint8_t x_771; -lean_dec(x_583); -x_770 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__14; -x_771 = lean_string_dec_eq(x_582, x_770); -lean_dec(x_582); -if (x_771 == 0) +lean_object* x_772; uint8_t x_773; +lean_dec(x_585); +x_772 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__14; +x_773 = lean_string_dec_eq(x_584, x_772); +lean_dec(x_584); +if (x_773 == 0) { -lean_object* x_772; lean_object* x_773; -lean_dec(x_580); +lean_object* x_774; lean_object* x_775; +lean_dec(x_582); +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_3); -x_772 = lean_unsigned_to_nat(0u); -lean_ctor_set(x_14, 1, x_772); -lean_ctor_set(x_14, 0, x_2); -x_773 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_773, 0, x_14); -lean_ctor_set(x_773, 1, x_13); -return x_773; +x_774 = lean_unsigned_to_nat(0u); +lean_ctor_set(x_16, 1, x_774); +lean_ctor_set(x_16, 0, x_2); +x_775 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_775, 0, x_16); +lean_ctor_set(x_775, 1, x_15); +return x_775; } else { -lean_object* x_774; lean_object* x_775; uint8_t x_776; -x_774 = lean_array_get_size(x_580); -x_775 = lean_unsigned_to_nat(4u); -x_776 = lean_nat_dec_eq(x_774, x_775); -lean_dec(x_774); -if (x_776 == 0) +lean_object* x_776; lean_object* x_777; uint8_t x_778; +x_776 = lean_array_get_size(x_582); +x_777 = lean_unsigned_to_nat(4u); +x_778 = lean_nat_dec_eq(x_776, x_777); +lean_dec(x_776); +if (x_778 == 0) { -lean_object* x_777; lean_object* x_778; -lean_dec(x_580); +lean_object* x_779; lean_object* x_780; +lean_dec(x_582); +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_3); -x_777 = lean_unsigned_to_nat(0u); -lean_ctor_set(x_14, 1, x_777); -lean_ctor_set(x_14, 0, x_2); -x_778 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_778, 0, x_14); -lean_ctor_set(x_778, 1, x_13); -return x_778; +x_779 = lean_unsigned_to_nat(0u); +lean_ctor_set(x_16, 1, x_779); +lean_ctor_set(x_16, 0, x_2); +x_780 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_780, 0, x_16); +lean_ctor_set(x_780, 1, x_15); +return x_780; } else { -lean_object* x_779; lean_object* x_780; -x_779 = lean_unsigned_to_nat(0u); -x_780 = lean_array_fget(x_580, x_779); -switch (lean_obj_tag(x_780)) { +lean_object* x_781; lean_object* x_782; +x_781 = lean_unsigned_to_nat(0u); +x_782 = lean_array_fget(x_582, x_781); +switch (lean_obj_tag(x_782)) { case 4: { -lean_object* x_781; -x_781 = lean_ctor_get(x_780, 0); -lean_inc(x_781); -if (lean_obj_tag(x_781) == 1) -{ -lean_object* x_782; -x_782 = lean_ctor_get(x_781, 0); -lean_inc(x_782); -if (lean_obj_tag(x_782) == 0) -{ -lean_object* x_783; lean_object* x_784; lean_object* x_785; uint8_t x_786; -x_783 = lean_ctor_get(x_780, 1); +lean_object* x_783; +x_783 = lean_ctor_get(x_782, 0); lean_inc(x_783); -lean_dec(x_780); -x_784 = lean_ctor_get(x_781, 1); +if (lean_obj_tag(x_783) == 1) +{ +lean_object* x_784; +x_784 = lean_ctor_get(x_783, 0); lean_inc(x_784); -lean_dec(x_781); -x_785 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__9; -x_786 = lean_string_dec_eq(x_784, x_785); -if (x_786 == 0) -{ -lean_object* x_787; uint8_t x_788; -x_787 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__1; -x_788 = lean_string_dec_eq(x_784, x_787); -lean_dec(x_784); -if (x_788 == 0) +if (lean_obj_tag(x_784) == 0) { -lean_object* x_789; +lean_object* x_785; lean_object* x_786; lean_object* x_787; uint8_t x_788; +x_785 = lean_ctor_get(x_782, 1); +lean_inc(x_785); +lean_dec(x_782); +x_786 = lean_ctor_get(x_783, 1); +lean_inc(x_786); lean_dec(x_783); -lean_dec(x_580); +x_787 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__9; +x_788 = lean_string_dec_eq(x_786, x_787); +if (x_788 == 0) +{ +lean_object* x_789; uint8_t x_790; +x_789 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__1; +x_790 = lean_string_dec_eq(x_786, x_789); +lean_dec(x_786); +if (x_790 == 0) +{ +lean_object* x_791; +lean_dec(x_785); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_779); -lean_ctor_set(x_14, 0, x_2); -x_789 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_789, 0, x_14); -lean_ctor_set(x_789, 1, x_13); -return x_789; +lean_ctor_set(x_16, 1, x_781); +lean_ctor_set(x_16, 0, x_2); +x_791 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_791, 0, x_16); +lean_ctor_set(x_791, 1, x_15); +return x_791; } else { -if (lean_obj_tag(x_783) == 0) +if (lean_obj_tag(x_785) == 0) { -lean_object* x_790; lean_object* x_791; lean_object* x_792; lean_object* x_793; lean_object* x_794; lean_object* x_795; lean_object* x_796; -lean_free_object(x_14); -x_790 = lean_unsigned_to_nat(2u); -x_791 = lean_array_fget(x_580, x_790); -x_792 = lean_unsigned_to_nat(3u); -x_793 = lean_array_fget(x_580, x_792); -lean_dec(x_580); -x_794 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__56; -x_795 = l_Lean_mkApp3(x_794, x_791, x_793, x_3); -x_796 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_795, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_796; +lean_object* x_792; lean_object* x_793; lean_object* x_794; lean_object* x_795; lean_object* x_796; lean_object* x_797; lean_object* x_798; +lean_free_object(x_16); +x_792 = lean_unsigned_to_nat(2u); +x_793 = lean_array_fget(x_582, x_792); +x_794 = lean_unsigned_to_nat(3u); +x_795 = lean_array_fget(x_582, x_794); +lean_dec(x_582); +x_796 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__56; +x_797 = l_Lean_mkApp3(x_796, x_793, x_795, x_3); +x_798 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_797, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_798; } else { -lean_object* x_797; -lean_dec(x_783); -lean_dec(x_580); +lean_object* x_799; +lean_dec(x_785); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_779); -lean_ctor_set(x_14, 0, x_2); -x_797 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_797, 0, x_14); -lean_ctor_set(x_797, 1, x_13); -return x_797; +lean_ctor_set(x_16, 1, x_781); +lean_ctor_set(x_16, 0, x_2); +x_799 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_799, 0, x_16); +lean_ctor_set(x_799, 1, x_15); +return x_799; } } } else { -lean_dec(x_784); -if (lean_obj_tag(x_783) == 0) -{ -lean_object* x_798; lean_object* x_799; lean_object* x_800; lean_object* x_801; lean_object* x_802; lean_object* x_803; lean_object* x_804; -lean_free_object(x_14); -x_798 = lean_unsigned_to_nat(2u); -x_799 = lean_array_fget(x_580, x_798); -x_800 = lean_unsigned_to_nat(3u); -x_801 = lean_array_fget(x_580, x_800); -lean_dec(x_580); -x_802 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__59; -x_803 = l_Lean_mkApp3(x_802, x_799, x_801, x_3); -x_804 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_803, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_804; +lean_dec(x_786); +if (lean_obj_tag(x_785) == 0) +{ +lean_object* x_800; lean_object* x_801; lean_object* x_802; lean_object* x_803; lean_object* x_804; lean_object* x_805; lean_object* x_806; +lean_free_object(x_16); +x_800 = lean_unsigned_to_nat(2u); +x_801 = lean_array_fget(x_582, x_800); +x_802 = lean_unsigned_to_nat(3u); +x_803 = lean_array_fget(x_582, x_802); +lean_dec(x_582); +x_804 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__59; +x_805 = l_Lean_mkApp3(x_804, x_801, x_803, x_3); +x_806 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_805, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_806; } else { -lean_object* x_805; -lean_dec(x_783); -lean_dec(x_580); +lean_object* x_807; +lean_dec(x_785); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_779); -lean_ctor_set(x_14, 0, x_2); -x_805 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_805, 0, x_14); -lean_ctor_set(x_805, 1, x_13); -return x_805; +lean_ctor_set(x_16, 1, x_781); +lean_ctor_set(x_16, 0, x_2); +x_807 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_807, 0, x_16); +lean_ctor_set(x_807, 1, x_15); +return x_807; } } } else { -lean_object* x_806; +lean_object* x_808; +lean_dec(x_784); +lean_dec(x_783); lean_dec(x_782); -lean_dec(x_781); -lean_dec(x_780); -lean_dec(x_580); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_779); -lean_ctor_set(x_14, 0, x_2); -x_806 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_806, 0, x_14); -lean_ctor_set(x_806, 1, x_13); -return x_806; +lean_ctor_set(x_16, 1, x_781); +lean_ctor_set(x_16, 0, x_2); +x_808 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_808, 0, x_16); +lean_ctor_set(x_808, 1, x_15); +return x_808; } } else { -lean_object* x_807; -lean_dec(x_781); -lean_dec(x_780); -lean_dec(x_580); +lean_object* x_809; +lean_dec(x_783); +lean_dec(x_782); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_779); -lean_ctor_set(x_14, 0, x_2); -x_807 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_807, 0, x_14); -lean_ctor_set(x_807, 1, x_13); -return x_807; +lean_ctor_set(x_16, 1, x_781); +lean_ctor_set(x_16, 0, x_2); +x_809 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_809, 0, x_16); +lean_ctor_set(x_809, 1, x_15); +return x_809; } } case 5: { -lean_object* x_808; -x_808 = lean_ctor_get(x_780, 0); -lean_inc(x_808); -if (lean_obj_tag(x_808) == 4) -{ -lean_object* x_809; -x_809 = lean_ctor_get(x_808, 0); -lean_inc(x_809); -if (lean_obj_tag(x_809) == 1) -{ lean_object* x_810; -x_810 = lean_ctor_get(x_809, 0); +x_810 = lean_ctor_get(x_782, 0); lean_inc(x_810); -if (lean_obj_tag(x_810) == 0) +if (lean_obj_tag(x_810) == 4) { -lean_object* x_811; lean_object* x_812; lean_object* x_813; lean_object* x_814; uint8_t x_815; -x_811 = lean_ctor_get(x_780, 1); +lean_object* x_811; +x_811 = lean_ctor_get(x_810, 0); lean_inc(x_811); -lean_dec(x_780); -x_812 = lean_ctor_get(x_808, 1); +if (lean_obj_tag(x_811) == 1) +{ +lean_object* x_812; +x_812 = lean_ctor_get(x_811, 0); lean_inc(x_812); -lean_dec(x_808); -x_813 = lean_ctor_get(x_809, 1); -lean_inc(x_813); -lean_dec(x_809); -x_814 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__4; -x_815 = lean_string_dec_eq(x_813, x_814); -lean_dec(x_813); -if (x_815 == 0) +if (lean_obj_tag(x_812) == 0) { -lean_object* x_816; -lean_dec(x_812); +lean_object* x_813; lean_object* x_814; lean_object* x_815; lean_object* x_816; uint8_t x_817; +x_813 = lean_ctor_get(x_782, 1); +lean_inc(x_813); +lean_dec(x_782); +x_814 = lean_ctor_get(x_810, 1); +lean_inc(x_814); +lean_dec(x_810); +x_815 = lean_ctor_get(x_811, 1); +lean_inc(x_815); lean_dec(x_811); -lean_dec(x_580); +x_816 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__4; +x_817 = lean_string_dec_eq(x_815, x_816); +lean_dec(x_815); +if (x_817 == 0) +{ +lean_object* x_818; +lean_dec(x_814); +lean_dec(x_813); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_779); -lean_ctor_set(x_14, 0, x_2); -x_816 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_816, 0, x_14); -lean_ctor_set(x_816, 1, x_13); -return x_816; +lean_ctor_set(x_16, 1, x_781); +lean_ctor_set(x_16, 0, x_2); +x_818 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_818, 0, x_16); +lean_ctor_set(x_818, 1, x_15); +return x_818; } else { -if (lean_obj_tag(x_812) == 0) +if (lean_obj_tag(x_814) == 0) { -lean_object* x_817; lean_object* x_818; lean_object* x_819; lean_object* x_820; lean_object* x_821; lean_object* x_822; lean_object* x_823; -lean_free_object(x_14); -x_817 = lean_unsigned_to_nat(2u); -x_818 = lean_array_fget(x_580, x_817); -x_819 = lean_unsigned_to_nat(3u); -x_820 = lean_array_fget(x_580, x_819); -lean_dec(x_580); -x_821 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__62; -x_822 = l_Lean_mkApp4(x_821, x_811, x_818, x_820, x_3); -x_823 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_822, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_823; +lean_object* x_819; lean_object* x_820; lean_object* x_821; lean_object* x_822; lean_object* x_823; lean_object* x_824; lean_object* x_825; +lean_free_object(x_16); +x_819 = lean_unsigned_to_nat(2u); +x_820 = lean_array_fget(x_582, x_819); +x_821 = lean_unsigned_to_nat(3u); +x_822 = lean_array_fget(x_582, x_821); +lean_dec(x_582); +x_823 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__62; +x_824 = l_Lean_mkApp4(x_823, x_813, x_820, x_822, x_3); +x_825 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_824, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_825; } else { -lean_object* x_824; -lean_dec(x_812); -lean_dec(x_811); -lean_dec(x_580); +lean_object* x_826; +lean_dec(x_814); +lean_dec(x_813); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_779); -lean_ctor_set(x_14, 0, x_2); -x_824 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_824, 0, x_14); -lean_ctor_set(x_824, 1, x_13); -return x_824; +lean_ctor_set(x_16, 1, x_781); +lean_ctor_set(x_16, 0, x_2); +x_826 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_826, 0, x_16); +lean_ctor_set(x_826, 1, x_15); +return x_826; } } } else { -lean_object* x_825; +lean_object* x_827; +lean_dec(x_812); +lean_dec(x_811); lean_dec(x_810); -lean_dec(x_809); -lean_dec(x_808); -lean_dec(x_780); -lean_dec(x_580); +lean_dec(x_782); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_779); -lean_ctor_set(x_14, 0, x_2); -x_825 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_825, 0, x_14); -lean_ctor_set(x_825, 1, x_13); -return x_825; +lean_ctor_set(x_16, 1, x_781); +lean_ctor_set(x_16, 0, x_2); +x_827 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_827, 0, x_16); +lean_ctor_set(x_827, 1, x_15); +return x_827; } } else { -lean_object* x_826; -lean_dec(x_809); -lean_dec(x_808); -lean_dec(x_780); -lean_dec(x_580); +lean_object* x_828; +lean_dec(x_811); +lean_dec(x_810); +lean_dec(x_782); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_779); -lean_ctor_set(x_14, 0, x_2); -x_826 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_826, 0, x_14); -lean_ctor_set(x_826, 1, x_13); -return x_826; +lean_ctor_set(x_16, 1, x_781); +lean_ctor_set(x_16, 0, x_2); +x_828 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_828, 0, x_16); +lean_ctor_set(x_828, 1, x_15); +return x_828; } } else { -lean_object* x_827; -lean_dec(x_808); -lean_dec(x_780); -lean_dec(x_580); +lean_object* x_829; +lean_dec(x_810); +lean_dec(x_782); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_779); -lean_ctor_set(x_14, 0, x_2); -x_827 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_827, 0, x_14); -lean_ctor_set(x_827, 1, x_13); -return x_827; +lean_ctor_set(x_16, 1, x_781); +lean_ctor_set(x_16, 0, x_2); +x_829 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_829, 0, x_16); +lean_ctor_set(x_829, 1, x_15); +return x_829; } } default: { -lean_object* x_828; -lean_dec(x_780); -lean_dec(x_580); +lean_object* x_830; +lean_dec(x_782); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_779); -lean_ctor_set(x_14, 0, x_2); -x_828 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_828, 0, x_14); -lean_ctor_set(x_828, 1, x_13); -return x_828; +lean_ctor_set(x_16, 1, x_781); +lean_ctor_set(x_16, 0, x_2); +x_830 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_830, 0, x_16); +lean_ctor_set(x_830, 1, x_15); +return x_830; } } } @@ -23813,246 +24210,250 @@ return x_828; } else { -lean_object* x_829; uint8_t x_830; -lean_dec(x_583); -x_829 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__11; -x_830 = lean_string_dec_eq(x_582, x_829); -lean_dec(x_582); -if (x_830 == 0) +lean_object* x_831; uint8_t x_832; +lean_dec(x_585); +x_831 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__11; +x_832 = lean_string_dec_eq(x_584, x_831); +lean_dec(x_584); +if (x_832 == 0) { -lean_object* x_831; lean_object* x_832; -lean_dec(x_580); +lean_object* x_833; lean_object* x_834; +lean_dec(x_582); +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_3); -x_831 = lean_unsigned_to_nat(0u); -lean_ctor_set(x_14, 1, x_831); -lean_ctor_set(x_14, 0, x_2); -x_832 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_832, 0, x_14); -lean_ctor_set(x_832, 1, x_13); -return x_832; +x_833 = lean_unsigned_to_nat(0u); +lean_ctor_set(x_16, 1, x_833); +lean_ctor_set(x_16, 0, x_2); +x_834 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_834, 0, x_16); +lean_ctor_set(x_834, 1, x_15); +return x_834; } else { -lean_object* x_833; lean_object* x_834; uint8_t x_835; -x_833 = lean_array_get_size(x_580); -x_834 = lean_unsigned_to_nat(4u); -x_835 = lean_nat_dec_eq(x_833, x_834); -lean_dec(x_833); -if (x_835 == 0) +lean_object* x_835; lean_object* x_836; uint8_t x_837; +x_835 = lean_array_get_size(x_582); +x_836 = lean_unsigned_to_nat(4u); +x_837 = lean_nat_dec_eq(x_835, x_836); +lean_dec(x_835); +if (x_837 == 0) { -lean_object* x_836; lean_object* x_837; -lean_dec(x_580); +lean_object* x_838; lean_object* x_839; +lean_dec(x_582); +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_3); -x_836 = lean_unsigned_to_nat(0u); -lean_ctor_set(x_14, 1, x_836); -lean_ctor_set(x_14, 0, x_2); -x_837 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_837, 0, x_14); -lean_ctor_set(x_837, 1, x_13); -return x_837; +x_838 = lean_unsigned_to_nat(0u); +lean_ctor_set(x_16, 1, x_838); +lean_ctor_set(x_16, 0, x_2); +x_839 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_839, 0, x_16); +lean_ctor_set(x_839, 1, x_15); +return x_839; } else { -lean_object* x_838; lean_object* x_839; -x_838 = lean_unsigned_to_nat(0u); -x_839 = lean_array_fget(x_580, x_838); -switch (lean_obj_tag(x_839)) { +lean_object* x_840; lean_object* x_841; +x_840 = lean_unsigned_to_nat(0u); +x_841 = lean_array_fget(x_582, x_840); +switch (lean_obj_tag(x_841)) { case 4: { -lean_object* x_840; -x_840 = lean_ctor_get(x_839, 0); -lean_inc(x_840); -if (lean_obj_tag(x_840) == 1) -{ -lean_object* x_841; -x_841 = lean_ctor_get(x_840, 0); -lean_inc(x_841); -if (lean_obj_tag(x_841) == 0) -{ -lean_object* x_842; lean_object* x_843; lean_object* x_844; uint8_t x_845; -x_842 = lean_ctor_get(x_839, 1); +lean_object* x_842; +x_842 = lean_ctor_get(x_841, 0); lean_inc(x_842); -lean_dec(x_839); -x_843 = lean_ctor_get(x_840, 1); +if (lean_obj_tag(x_842) == 1) +{ +lean_object* x_843; +x_843 = lean_ctor_get(x_842, 0); lean_inc(x_843); -lean_dec(x_840); -x_844 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__9; -x_845 = lean_string_dec_eq(x_843, x_844); -if (x_845 == 0) -{ -lean_object* x_846; uint8_t x_847; -x_846 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__1; -x_847 = lean_string_dec_eq(x_843, x_846); -lean_dec(x_843); -if (x_847 == 0) +if (lean_obj_tag(x_843) == 0) { -lean_object* x_848; +lean_object* x_844; lean_object* x_845; lean_object* x_846; uint8_t x_847; +x_844 = lean_ctor_get(x_841, 1); +lean_inc(x_844); +lean_dec(x_841); +x_845 = lean_ctor_get(x_842, 1); +lean_inc(x_845); lean_dec(x_842); -lean_dec(x_580); +x_846 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__9; +x_847 = lean_string_dec_eq(x_845, x_846); +if (x_847 == 0) +{ +lean_object* x_848; uint8_t x_849; +x_848 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__1; +x_849 = lean_string_dec_eq(x_845, x_848); +lean_dec(x_845); +if (x_849 == 0) +{ +lean_object* x_850; +lean_dec(x_844); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_838); -lean_ctor_set(x_14, 0, x_2); -x_848 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_848, 0, x_14); -lean_ctor_set(x_848, 1, x_13); -return x_848; +lean_ctor_set(x_16, 1, x_840); +lean_ctor_set(x_16, 0, x_2); +x_850 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_850, 0, x_16); +lean_ctor_set(x_850, 1, x_15); +return x_850; } else { -if (lean_obj_tag(x_842) == 0) +if (lean_obj_tag(x_844) == 0) { -lean_object* x_849; lean_object* x_850; lean_object* x_851; lean_object* x_852; lean_object* x_853; lean_object* x_854; lean_object* x_855; -lean_free_object(x_14); -x_849 = lean_unsigned_to_nat(2u); -x_850 = lean_array_fget(x_580, x_849); -x_851 = lean_unsigned_to_nat(3u); -x_852 = lean_array_fget(x_580, x_851); -lean_dec(x_580); -x_853 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__65; -x_854 = l_Lean_mkApp3(x_853, x_850, x_852, x_3); -x_855 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_854, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_855; +lean_object* x_851; lean_object* x_852; lean_object* x_853; lean_object* x_854; lean_object* x_855; lean_object* x_856; lean_object* x_857; +lean_free_object(x_16); +x_851 = lean_unsigned_to_nat(2u); +x_852 = lean_array_fget(x_582, x_851); +x_853 = lean_unsigned_to_nat(3u); +x_854 = lean_array_fget(x_582, x_853); +lean_dec(x_582); +x_855 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__65; +x_856 = l_Lean_mkApp3(x_855, x_852, x_854, x_3); +x_857 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_856, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_857; } else { -lean_object* x_856; -lean_dec(x_842); -lean_dec(x_580); +lean_object* x_858; +lean_dec(x_844); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_838); -lean_ctor_set(x_14, 0, x_2); -x_856 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_856, 0, x_14); -lean_ctor_set(x_856, 1, x_13); -return x_856; +lean_ctor_set(x_16, 1, x_840); +lean_ctor_set(x_16, 0, x_2); +x_858 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_858, 0, x_16); +lean_ctor_set(x_858, 1, x_15); +return x_858; } } } else { -lean_dec(x_843); -if (lean_obj_tag(x_842) == 0) +lean_dec(x_845); +if (lean_obj_tag(x_844) == 0) { -lean_object* x_857; lean_object* x_858; lean_object* x_859; lean_object* x_860; lean_object* x_861; -x_857 = lean_unsigned_to_nat(2u); -x_858 = lean_array_fget(x_580, x_857); -x_859 = lean_unsigned_to_nat(3u); -x_860 = lean_array_fget(x_580, x_859); -lean_dec(x_580); -lean_inc(x_858); -x_861 = l_Lean_Expr_int_x3f(x_858); -if (lean_obj_tag(x_861) == 0) +lean_object* x_859; lean_object* x_860; lean_object* x_861; lean_object* x_862; lean_object* x_863; +x_859 = lean_unsigned_to_nat(2u); +x_860 = lean_array_fget(x_582, x_859); +x_861 = lean_unsigned_to_nat(3u); +x_862 = lean_array_fget(x_582, x_861); +lean_dec(x_582); +lean_inc(x_860); +x_863 = l_Lean_Expr_int_x3f(x_860); +if (lean_obj_tag(x_863) == 0) { -lean_object* x_862; lean_object* x_863; lean_object* x_864; -lean_free_object(x_14); -x_862 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__68; -x_863 = l_Lean_mkApp3(x_862, x_860, x_858, x_3); -x_864 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_863, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_864; +lean_object* x_864; lean_object* x_865; lean_object* x_866; +lean_free_object(x_16); +x_864 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__68; +x_865 = l_Lean_mkApp3(x_864, x_862, x_860, x_3); +x_866 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_865, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_866; } else { -lean_object* x_865; lean_object* x_866; uint8_t x_867; -x_865 = lean_ctor_get(x_861, 0); -lean_inc(x_865); -lean_dec(x_861); -x_866 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__8; -x_867 = lean_int_dec_eq(x_865, x_866); -lean_dec(x_865); -if (x_867 == 0) +lean_object* x_867; lean_object* x_868; uint8_t x_869; +x_867 = lean_ctor_get(x_863, 0); +lean_inc(x_867); +lean_dec(x_863); +x_868 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__8; +x_869 = lean_int_dec_eq(x_867, x_868); +lean_dec(x_867); +if (x_869 == 0) { -lean_object* x_868; lean_object* x_869; lean_object* x_870; -lean_free_object(x_14); -x_868 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__68; -x_869 = l_Lean_mkApp3(x_868, x_860, x_858, x_3); -x_870 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_869, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_870; +lean_object* x_870; lean_object* x_871; lean_object* x_872; +lean_free_object(x_16); +x_870 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__68; +x_871 = l_Lean_mkApp3(x_870, x_862, x_860, x_3); +x_872 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_871, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_872; } else { -lean_object* x_871; -lean_dec(x_858); -x_871 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality(x_2, x_3, x_860, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_871) == 0) +lean_object* x_873; +lean_dec(x_860); +x_873 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality(x_2, x_3, x_862, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_873) == 0) { -uint8_t x_872; -x_872 = !lean_is_exclusive(x_871); -if (x_872 == 0) +uint8_t x_874; +x_874 = !lean_is_exclusive(x_873); +if (x_874 == 0) { -lean_object* x_873; lean_object* x_874; -x_873 = lean_ctor_get(x_871, 0); -x_874 = lean_unsigned_to_nat(1u); -lean_ctor_set(x_14, 1, x_874); -lean_ctor_set(x_14, 0, x_873); -lean_ctor_set(x_871, 0, x_14); -return x_871; +lean_object* x_875; lean_object* x_876; +x_875 = lean_ctor_get(x_873, 0); +x_876 = lean_unsigned_to_nat(1u); +lean_ctor_set(x_16, 1, x_876); +lean_ctor_set(x_16, 0, x_875); +lean_ctor_set(x_873, 0, x_16); +return x_873; } else { -lean_object* x_875; lean_object* x_876; lean_object* x_877; lean_object* x_878; -x_875 = lean_ctor_get(x_871, 0); -x_876 = lean_ctor_get(x_871, 1); -lean_inc(x_876); -lean_inc(x_875); -lean_dec(x_871); -x_877 = lean_unsigned_to_nat(1u); -lean_ctor_set(x_14, 1, x_877); -lean_ctor_set(x_14, 0, x_875); -x_878 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_878, 0, x_14); -lean_ctor_set(x_878, 1, x_876); -return x_878; +lean_object* x_877; lean_object* x_878; lean_object* x_879; lean_object* x_880; +x_877 = lean_ctor_get(x_873, 0); +x_878 = lean_ctor_get(x_873, 1); +lean_inc(x_878); +lean_inc(x_877); +lean_dec(x_873); +x_879 = lean_unsigned_to_nat(1u); +lean_ctor_set(x_16, 1, x_879); +lean_ctor_set(x_16, 0, x_877); +x_880 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_880, 0, x_16); +lean_ctor_set(x_880, 1, x_878); +return x_880; } } else { -uint8_t x_879; -lean_free_object(x_14); -x_879 = !lean_is_exclusive(x_871); -if (x_879 == 0) +uint8_t x_881; +lean_free_object(x_16); +x_881 = !lean_is_exclusive(x_873); +if (x_881 == 0) { -return x_871; +return x_873; } else { -lean_object* x_880; lean_object* x_881; lean_object* x_882; -x_880 = lean_ctor_get(x_871, 0); -x_881 = lean_ctor_get(x_871, 1); -lean_inc(x_881); -lean_inc(x_880); -lean_dec(x_871); -x_882 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_882, 0, x_880); -lean_ctor_set(x_882, 1, x_881); -return x_882; +lean_object* x_882; lean_object* x_883; lean_object* x_884; +x_882 = lean_ctor_get(x_873, 0); +x_883 = lean_ctor_get(x_873, 1); +lean_inc(x_883); +lean_inc(x_882); +lean_dec(x_873); +x_884 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_884, 0, x_882); +lean_ctor_set(x_884, 1, x_883); +return x_884; } } } @@ -24060,249 +24461,258 @@ return x_882; } else { -lean_object* x_883; -lean_dec(x_842); -lean_dec(x_580); +lean_object* x_885; +lean_dec(x_844); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_838); -lean_ctor_set(x_14, 0, x_2); -x_883 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_883, 0, x_14); -lean_ctor_set(x_883, 1, x_13); -return x_883; +lean_ctor_set(x_16, 1, x_840); +lean_ctor_set(x_16, 0, x_2); +x_885 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_885, 0, x_16); +lean_ctor_set(x_885, 1, x_15); +return x_885; } } } else { -lean_object* x_884; +lean_object* x_886; +lean_dec(x_843); +lean_dec(x_842); lean_dec(x_841); -lean_dec(x_840); -lean_dec(x_839); -lean_dec(x_580); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_838); -lean_ctor_set(x_14, 0, x_2); -x_884 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_884, 0, x_14); -lean_ctor_set(x_884, 1, x_13); -return x_884; +lean_ctor_set(x_16, 1, x_840); +lean_ctor_set(x_16, 0, x_2); +x_886 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_886, 0, x_16); +lean_ctor_set(x_886, 1, x_15); +return x_886; } } else { -lean_object* x_885; -lean_dec(x_840); -lean_dec(x_839); -lean_dec(x_580); +lean_object* x_887; +lean_dec(x_842); +lean_dec(x_841); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_838); -lean_ctor_set(x_14, 0, x_2); -x_885 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_885, 0, x_14); -lean_ctor_set(x_885, 1, x_13); -return x_885; +lean_ctor_set(x_16, 1, x_840); +lean_ctor_set(x_16, 0, x_2); +x_887 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_887, 0, x_16); +lean_ctor_set(x_887, 1, x_15); +return x_887; } } case 5: { -lean_object* x_886; -x_886 = lean_ctor_get(x_839, 0); -lean_inc(x_886); -if (lean_obj_tag(x_886) == 4) -{ -lean_object* x_887; -x_887 = lean_ctor_get(x_886, 0); -lean_inc(x_887); -if (lean_obj_tag(x_887) == 1) -{ lean_object* x_888; -x_888 = lean_ctor_get(x_887, 0); +x_888 = lean_ctor_get(x_841, 0); lean_inc(x_888); -if (lean_obj_tag(x_888) == 0) +if (lean_obj_tag(x_888) == 4) { -lean_object* x_889; lean_object* x_890; lean_object* x_891; lean_object* x_892; uint8_t x_893; -x_889 = lean_ctor_get(x_839, 1); +lean_object* x_889; +x_889 = lean_ctor_get(x_888, 0); lean_inc(x_889); -lean_dec(x_839); -x_890 = lean_ctor_get(x_886, 1); +if (lean_obj_tag(x_889) == 1) +{ +lean_object* x_890; +x_890 = lean_ctor_get(x_889, 0); lean_inc(x_890); -lean_dec(x_886); -x_891 = lean_ctor_get(x_887, 1); -lean_inc(x_891); -lean_dec(x_887); -x_892 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__4; -x_893 = lean_string_dec_eq(x_891, x_892); -lean_dec(x_891); -if (x_893 == 0) +if (lean_obj_tag(x_890) == 0) { -lean_object* x_894; -lean_dec(x_890); +lean_object* x_891; lean_object* x_892; lean_object* x_893; lean_object* x_894; uint8_t x_895; +x_891 = lean_ctor_get(x_841, 1); +lean_inc(x_891); +lean_dec(x_841); +x_892 = lean_ctor_get(x_888, 1); +lean_inc(x_892); +lean_dec(x_888); +x_893 = lean_ctor_get(x_889, 1); +lean_inc(x_893); lean_dec(x_889); -lean_dec(x_580); +x_894 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__4; +x_895 = lean_string_dec_eq(x_893, x_894); +lean_dec(x_893); +if (x_895 == 0) +{ +lean_object* x_896; +lean_dec(x_892); +lean_dec(x_891); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_838); -lean_ctor_set(x_14, 0, x_2); -x_894 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_894, 0, x_14); -lean_ctor_set(x_894, 1, x_13); -return x_894; +lean_ctor_set(x_16, 1, x_840); +lean_ctor_set(x_16, 0, x_2); +x_896 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_896, 0, x_16); +lean_ctor_set(x_896, 1, x_15); +return x_896; } else { -if (lean_obj_tag(x_890) == 0) +if (lean_obj_tag(x_892) == 0) { -lean_object* x_895; lean_object* x_896; lean_object* x_897; lean_object* x_898; lean_object* x_899; lean_object* x_900; lean_object* x_901; -lean_free_object(x_14); -x_895 = lean_unsigned_to_nat(2u); -x_896 = lean_array_fget(x_580, x_895); -x_897 = lean_unsigned_to_nat(3u); -x_898 = lean_array_fget(x_580, x_897); -lean_dec(x_580); -x_899 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__71; -x_900 = l_Lean_mkApp4(x_899, x_889, x_896, x_898, x_3); -x_901 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_900, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_901; +lean_object* x_897; lean_object* x_898; lean_object* x_899; lean_object* x_900; lean_object* x_901; lean_object* x_902; lean_object* x_903; +lean_free_object(x_16); +x_897 = lean_unsigned_to_nat(2u); +x_898 = lean_array_fget(x_582, x_897); +x_899 = lean_unsigned_to_nat(3u); +x_900 = lean_array_fget(x_582, x_899); +lean_dec(x_582); +x_901 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__71; +x_902 = l_Lean_mkApp4(x_901, x_891, x_898, x_900, x_3); +x_903 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_902, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_903; } else { -lean_object* x_902; -lean_dec(x_890); -lean_dec(x_889); -lean_dec(x_580); +lean_object* x_904; +lean_dec(x_892); +lean_dec(x_891); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_838); -lean_ctor_set(x_14, 0, x_2); -x_902 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_902, 0, x_14); -lean_ctor_set(x_902, 1, x_13); -return x_902; +lean_ctor_set(x_16, 1, x_840); +lean_ctor_set(x_16, 0, x_2); +x_904 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_904, 0, x_16); +lean_ctor_set(x_904, 1, x_15); +return x_904; } } } else { -lean_object* x_903; +lean_object* x_905; +lean_dec(x_890); +lean_dec(x_889); lean_dec(x_888); -lean_dec(x_887); -lean_dec(x_886); -lean_dec(x_839); -lean_dec(x_580); +lean_dec(x_841); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_838); -lean_ctor_set(x_14, 0, x_2); -x_903 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_903, 0, x_14); -lean_ctor_set(x_903, 1, x_13); -return x_903; +lean_ctor_set(x_16, 1, x_840); +lean_ctor_set(x_16, 0, x_2); +x_905 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_905, 0, x_16); +lean_ctor_set(x_905, 1, x_15); +return x_905; } } else { -lean_object* x_904; -lean_dec(x_887); -lean_dec(x_886); -lean_dec(x_839); -lean_dec(x_580); +lean_object* x_906; +lean_dec(x_889); +lean_dec(x_888); +lean_dec(x_841); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_838); -lean_ctor_set(x_14, 0, x_2); -x_904 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_904, 0, x_14); -lean_ctor_set(x_904, 1, x_13); -return x_904; +lean_ctor_set(x_16, 1, x_840); +lean_ctor_set(x_16, 0, x_2); +x_906 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_906, 0, x_16); +lean_ctor_set(x_906, 1, x_15); +return x_906; } } else { -lean_object* x_905; -lean_dec(x_886); -lean_dec(x_839); -lean_dec(x_580); +lean_object* x_907; +lean_dec(x_888); +lean_dec(x_841); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_838); -lean_ctor_set(x_14, 0, x_2); -x_905 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_905, 0, x_14); -lean_ctor_set(x_905, 1, x_13); -return x_905; +lean_ctor_set(x_16, 1, x_840); +lean_ctor_set(x_16, 0, x_2); +x_907 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_907, 0, x_16); +lean_ctor_set(x_907, 1, x_15); +return x_907; } } default: { -lean_object* x_906; -lean_dec(x_839); -lean_dec(x_580); +lean_object* x_908; +lean_dec(x_841); +lean_dec(x_582); +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_3); -lean_ctor_set(x_14, 1, x_838); -lean_ctor_set(x_14, 0, x_2); -x_906 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_906, 0, x_14); -lean_ctor_set(x_906, 1, x_13); -return x_906; +lean_ctor_set(x_16, 1, x_840); +lean_ctor_set(x_16, 0, x_2); +x_908 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_908, 0, x_16); +lean_ctor_set(x_908, 1, x_15); +return x_908; } } } @@ -24311,2178 +24721,2241 @@ return x_906; } else { -lean_object* x_907; lean_object* x_908; lean_object* x_909; lean_object* x_910; uint8_t x_911; -x_907 = lean_ctor_get(x_14, 1); -lean_inc(x_907); -lean_dec(x_14); -x_908 = lean_ctor_get(x_15, 1); -lean_inc(x_908); -lean_dec(x_15); +lean_object* x_909; lean_object* x_910; lean_object* x_911; lean_object* x_912; uint8_t x_913; x_909 = lean_ctor_get(x_16, 1); lean_inc(x_909); lean_dec(x_16); -x_910 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__10; -x_911 = lean_string_dec_eq(x_909, x_910); -if (x_911 == 0) -{ -lean_object* x_912; uint8_t x_913; -x_912 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__13; -x_913 = lean_string_dec_eq(x_909, x_912); +x_910 = lean_ctor_get(x_17, 1); +lean_inc(x_910); +lean_dec(x_17); +x_911 = lean_ctor_get(x_18, 1); +lean_inc(x_911); +lean_dec(x_18); +x_912 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__10; +x_913 = lean_string_dec_eq(x_911, x_912); if (x_913 == 0) { lean_object* x_914; uint8_t x_915; -x_914 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__26; -x_915 = lean_string_dec_eq(x_909, x_914); +x_914 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__13; +x_915 = lean_string_dec_eq(x_911, x_914); if (x_915 == 0) { lean_object* x_916; uint8_t x_917; -x_916 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__27; -x_917 = lean_string_dec_eq(x_909, x_916); +x_916 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__26; +x_917 = lean_string_dec_eq(x_911, x_916); if (x_917 == 0) { lean_object* x_918; uint8_t x_919; -x_918 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__8; -x_919 = lean_string_dec_eq(x_909, x_918); +x_918 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__27; +x_919 = lean_string_dec_eq(x_911, x_918); if (x_919 == 0) { lean_object* x_920; uint8_t x_921; -x_920 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__7; -x_921 = lean_string_dec_eq(x_909, x_920); -lean_dec(x_909); +x_920 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__8; +x_921 = lean_string_dec_eq(x_911, x_920); if (x_921 == 0) { -lean_object* x_922; lean_object* x_923; lean_object* x_924; -lean_dec(x_908); -lean_dec(x_907); +lean_object* x_922; uint8_t x_923; +x_922 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__7; +x_923 = lean_string_dec_eq(x_911, x_922); +lean_dec(x_911); +if (x_923 == 0) +{ +lean_object* x_924; lean_object* x_925; lean_object* x_926; +lean_dec(x_910); +lean_dec(x_909); +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_3); -x_922 = lean_unsigned_to_nat(0u); -x_923 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_923, 0, x_2); -lean_ctor_set(x_923, 1, x_922); -x_924 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_924, 0, x_923); -lean_ctor_set(x_924, 1, x_13); -return x_924; +x_924 = lean_unsigned_to_nat(0u); +x_925 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_925, 0, x_2); +lean_ctor_set(x_925, 1, x_924); +x_926 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_926, 0, x_925); +lean_ctor_set(x_926, 1, x_15); +return x_926; } else { -lean_object* x_925; uint8_t x_926; -x_925 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__8; -x_926 = lean_string_dec_eq(x_908, x_925); -lean_dec(x_908); -if (x_926 == 0) +lean_object* x_927; uint8_t x_928; +x_927 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__8; +x_928 = lean_string_dec_eq(x_910, x_927); +lean_dec(x_910); +if (x_928 == 0) { -lean_object* x_927; lean_object* x_928; lean_object* x_929; -lean_dec(x_907); +lean_object* x_929; lean_object* x_930; lean_object* x_931; +lean_dec(x_909); +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_3); -x_927 = lean_unsigned_to_nat(0u); -x_928 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_928, 0, x_2); -lean_ctor_set(x_928, 1, x_927); -x_929 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_929, 0, x_928); -lean_ctor_set(x_929, 1, x_13); -return x_929; +x_929 = lean_unsigned_to_nat(0u); +x_930 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_930, 0, x_2); +lean_ctor_set(x_930, 1, x_929); +x_931 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_931, 0, x_930); +lean_ctor_set(x_931, 1, x_15); +return x_931; } else { -lean_object* x_930; lean_object* x_931; uint8_t x_932; -x_930 = lean_array_get_size(x_907); -x_931 = lean_unsigned_to_nat(4u); -x_932 = lean_nat_dec_eq(x_930, x_931); -lean_dec(x_930); -if (x_932 == 0) +lean_object* x_932; lean_object* x_933; uint8_t x_934; +x_932 = lean_array_get_size(x_909); +x_933 = lean_unsigned_to_nat(4u); +x_934 = lean_nat_dec_eq(x_932, x_933); +lean_dec(x_932); +if (x_934 == 0) { -lean_object* x_933; lean_object* x_934; lean_object* x_935; -lean_dec(x_907); +lean_object* x_935; lean_object* x_936; lean_object* x_937; +lean_dec(x_909); +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_3); -x_933 = lean_unsigned_to_nat(0u); -x_934 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_934, 0, x_2); -lean_ctor_set(x_934, 1, x_933); -x_935 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_935, 0, x_934); -lean_ctor_set(x_935, 1, x_13); -return x_935; +x_935 = lean_unsigned_to_nat(0u); +x_936 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_936, 0, x_2); +lean_ctor_set(x_936, 1, x_935); +x_937 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_937, 0, x_936); +lean_ctor_set(x_937, 1, x_15); +return x_937; } else { -lean_object* x_936; lean_object* x_937; -x_936 = lean_unsigned_to_nat(0u); -x_937 = lean_array_fget(x_907, x_936); -if (lean_obj_tag(x_937) == 4) -{ -lean_object* x_938; -x_938 = lean_ctor_get(x_937, 0); -lean_inc(x_938); -if (lean_obj_tag(x_938) == 1) -{ -lean_object* x_939; -x_939 = lean_ctor_get(x_938, 0); -lean_inc(x_939); -if (lean_obj_tag(x_939) == 0) +lean_object* x_938; lean_object* x_939; +x_938 = lean_unsigned_to_nat(0u); +x_939 = lean_array_fget(x_909, x_938); +if (lean_obj_tag(x_939) == 4) { -lean_object* x_940; lean_object* x_941; lean_object* x_942; uint8_t x_943; -x_940 = lean_ctor_get(x_937, 1); +lean_object* x_940; +x_940 = lean_ctor_get(x_939, 0); lean_inc(x_940); -lean_dec(x_937); -x_941 = lean_ctor_get(x_938, 1); +if (lean_obj_tag(x_940) == 1) +{ +lean_object* x_941; +x_941 = lean_ctor_get(x_940, 0); lean_inc(x_941); -lean_dec(x_938); -x_942 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__1; -x_943 = lean_string_dec_eq(x_941, x_942); -if (x_943 == 0) -{ -lean_object* x_944; uint8_t x_945; -x_944 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__9; -x_945 = lean_string_dec_eq(x_941, x_944); -lean_dec(x_941); -if (x_945 == 0) +if (lean_obj_tag(x_941) == 0) { -lean_object* x_946; lean_object* x_947; +lean_object* x_942; lean_object* x_943; lean_object* x_944; uint8_t x_945; +x_942 = lean_ctor_get(x_939, 1); +lean_inc(x_942); +lean_dec(x_939); +x_943 = lean_ctor_get(x_940, 1); +lean_inc(x_943); lean_dec(x_940); -lean_dec(x_907); -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_3); -x_946 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_946, 0, x_2); -lean_ctor_set(x_946, 1, x_936); -x_947 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_947, 0, x_946); -lean_ctor_set(x_947, 1, x_13); -return x_947; -} -else -{ -if (lean_obj_tag(x_940) == 0) +x_944 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__1; +x_945 = lean_string_dec_eq(x_943, x_944); +if (x_945 == 0) { -lean_object* x_948; lean_object* x_949; lean_object* x_950; lean_object* x_951; lean_object* x_952; lean_object* x_953; lean_object* x_954; -x_948 = lean_unsigned_to_nat(2u); -x_949 = lean_array_fget(x_907, x_948); -x_950 = lean_unsigned_to_nat(3u); -x_951 = lean_array_fget(x_907, x_950); -lean_dec(x_907); -x_952 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__30; -x_953 = l_Lean_mkApp3(x_952, x_949, x_951, x_3); -x_954 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_953, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_954; -} -else +lean_object* x_946; uint8_t x_947; +x_946 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__9; +x_947 = lean_string_dec_eq(x_943, x_946); +lean_dec(x_943); +if (x_947 == 0) { -lean_object* x_955; lean_object* x_956; -lean_dec(x_940); -lean_dec(x_907); +lean_object* x_948; lean_object* x_949; +lean_dec(x_942); +lean_dec(x_909); +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_3); -x_955 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_955, 0, x_2); -lean_ctor_set(x_955, 1, x_936); -x_956 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_956, 0, x_955); -lean_ctor_set(x_956, 1, x_13); -return x_956; -} -} +x_948 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_948, 0, x_2); +lean_ctor_set(x_948, 1, x_938); +x_949 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_949, 0, x_948); +lean_ctor_set(x_949, 1, x_15); +return x_949; } else { -lean_dec(x_941); -if (lean_obj_tag(x_940) == 0) +if (lean_obj_tag(x_942) == 0) { -lean_object* x_957; lean_object* x_958; lean_object* x_959; lean_object* x_960; lean_object* x_961; lean_object* x_962; lean_object* x_963; -x_957 = lean_unsigned_to_nat(2u); -x_958 = lean_array_fget(x_907, x_957); -x_959 = lean_unsigned_to_nat(3u); -x_960 = lean_array_fget(x_907, x_959); -lean_dec(x_907); -x_961 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__33; -x_962 = l_Lean_mkApp3(x_961, x_958, x_960, x_3); -x_963 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_962, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_963; +lean_object* x_950; lean_object* x_951; lean_object* x_952; lean_object* x_953; lean_object* x_954; lean_object* x_955; lean_object* x_956; +x_950 = lean_unsigned_to_nat(2u); +x_951 = lean_array_fget(x_909, x_950); +x_952 = lean_unsigned_to_nat(3u); +x_953 = lean_array_fget(x_909, x_952); +lean_dec(x_909); +x_954 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__30; +x_955 = l_Lean_mkApp3(x_954, x_951, x_953, x_3); +x_956 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_955, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_956; } else { -lean_object* x_964; lean_object* x_965; -lean_dec(x_940); -lean_dec(x_907); +lean_object* x_957; lean_object* x_958; +lean_dec(x_942); +lean_dec(x_909); +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_3); -x_964 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_964, 0, x_2); -lean_ctor_set(x_964, 1, x_936); -x_965 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_965, 0, x_964); -lean_ctor_set(x_965, 1, x_13); -return x_965; +x_957 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_957, 0, x_2); +lean_ctor_set(x_957, 1, x_938); +x_958 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_958, 0, x_957); +lean_ctor_set(x_958, 1, x_15); +return x_958; +} } } +else +{ +lean_dec(x_943); +if (lean_obj_tag(x_942) == 0) +{ +lean_object* x_959; lean_object* x_960; lean_object* x_961; lean_object* x_962; lean_object* x_963; lean_object* x_964; lean_object* x_965; +x_959 = lean_unsigned_to_nat(2u); +x_960 = lean_array_fget(x_909, x_959); +x_961 = lean_unsigned_to_nat(3u); +x_962 = lean_array_fget(x_909, x_961); +lean_dec(x_909); +x_963 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__33; +x_964 = l_Lean_mkApp3(x_963, x_960, x_962, x_3); +x_965 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_964, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_965; } else { lean_object* x_966; lean_object* x_967; -lean_dec(x_939); -lean_dec(x_938); -lean_dec(x_937); -lean_dec(x_907); +lean_dec(x_942); +lean_dec(x_909); +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_3); x_966 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_966, 0, x_2); -lean_ctor_set(x_966, 1, x_936); +lean_ctor_set(x_966, 1, x_938); x_967 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_967, 0, x_966); -lean_ctor_set(x_967, 1, x_13); +lean_ctor_set(x_967, 1, x_15); return x_967; } } +} else { lean_object* x_968; lean_object* x_969; -lean_dec(x_938); -lean_dec(x_937); -lean_dec(x_907); +lean_dec(x_941); +lean_dec(x_940); +lean_dec(x_939); +lean_dec(x_909); +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_3); x_968 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_968, 0, x_2); -lean_ctor_set(x_968, 1, x_936); +lean_ctor_set(x_968, 1, x_938); x_969 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_969, 0, x_968); -lean_ctor_set(x_969, 1, x_13); +lean_ctor_set(x_969, 1, x_15); return x_969; } } else { lean_object* x_970; lean_object* x_971; -lean_dec(x_937); -lean_dec(x_907); +lean_dec(x_940); +lean_dec(x_939); +lean_dec(x_909); +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_3); x_970 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_970, 0, x_2); -lean_ctor_set(x_970, 1, x_936); +lean_ctor_set(x_970, 1, x_938); x_971 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_971, 0, x_970); -lean_ctor_set(x_971, 1, x_13); +lean_ctor_set(x_971, 1, x_15); return x_971; } } +else +{ +lean_object* x_972; lean_object* x_973; +lean_dec(x_939); +lean_dec(x_909); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_3); +x_972 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_972, 0, x_2); +lean_ctor_set(x_972, 1, x_938); +x_973 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_973, 0, x_972); +lean_ctor_set(x_973, 1, x_15); +return x_973; +} +} } } } else { -lean_object* x_972; uint8_t x_973; +lean_object* x_974; uint8_t x_975; +lean_dec(x_911); lean_dec(x_909); -lean_dec(x_907); -x_972 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__16; -x_973 = lean_string_dec_eq(x_908, x_972); -lean_dec(x_908); -if (x_973 == 0) +x_974 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__16; +x_975 = lean_string_dec_eq(x_910, x_974); +lean_dec(x_910); +if (x_975 == 0) { -lean_object* x_974; lean_object* x_975; lean_object* x_976; +lean_object* x_976; lean_object* x_977; lean_object* x_978; +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_3); -x_974 = lean_unsigned_to_nat(0u); -x_975 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_975, 0, x_2); -lean_ctor_set(x_975, 1, x_974); -x_976 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_976, 0, x_975); -lean_ctor_set(x_976, 1, x_13); -return x_976; +x_976 = lean_unsigned_to_nat(0u); +x_977 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_977, 0, x_2); +lean_ctor_set(x_977, 1, x_976); +x_978 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_978, 0, x_977); +lean_ctor_set(x_978, 1, x_15); +return x_978; } else { -lean_object* x_977; lean_object* x_978; lean_object* x_979; lean_object* x_980; -x_977 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__5___closed__6; -x_978 = lean_array_push(x_977, x_3); -x_979 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__35; +lean_object* x_979; lean_object* x_980; lean_object* x_981; lean_object* x_982; +x_979 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__5___closed__6; +x_980 = lean_array_push(x_979, x_3); +x_981 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__35; +lean_inc(x_14); +lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -x_980 = l_Lean_Meta_mkAppM(x_979, x_978, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_980) == 0) +x_982 = l_Lean_Meta_mkAppM(x_981, x_980, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_982) == 0) { -lean_object* x_981; lean_object* x_982; lean_object* x_983; -x_981 = lean_ctor_get(x_980, 0); -lean_inc(x_981); -x_982 = lean_ctor_get(x_980, 1); -lean_inc(x_982); -lean_dec(x_980); -x_983 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_981, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_982); -return x_983; +lean_object* x_983; lean_object* x_984; lean_object* x_985; +x_983 = lean_ctor_get(x_982, 0); +lean_inc(x_983); +x_984 = lean_ctor_get(x_982, 1); +lean_inc(x_984); +lean_dec(x_982); +x_985 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_983, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_984); +return x_985; } else { -lean_object* x_984; lean_object* x_985; lean_object* x_986; lean_object* x_987; +lean_object* x_986; lean_object* x_987; lean_object* x_988; lean_object* x_989; +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_2); -x_984 = lean_ctor_get(x_980, 0); -lean_inc(x_984); -x_985 = lean_ctor_get(x_980, 1); -lean_inc(x_985); -if (lean_is_exclusive(x_980)) { - lean_ctor_release(x_980, 0); - lean_ctor_release(x_980, 1); - x_986 = x_980; +x_986 = lean_ctor_get(x_982, 0); +lean_inc(x_986); +x_987 = lean_ctor_get(x_982, 1); +lean_inc(x_987); +if (lean_is_exclusive(x_982)) { + lean_ctor_release(x_982, 0); + lean_ctor_release(x_982, 1); + x_988 = x_982; } else { - lean_dec_ref(x_980); - x_986 = lean_box(0); + lean_dec_ref(x_982); + x_988 = lean_box(0); } -if (lean_is_scalar(x_986)) { - x_987 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_988)) { + x_989 = lean_alloc_ctor(1, 2, 0); } else { - x_987 = x_986; + x_989 = x_988; } -lean_ctor_set(x_987, 0, x_984); -lean_ctor_set(x_987, 1, x_985); -return x_987; +lean_ctor_set(x_989, 0, x_986); +lean_ctor_set(x_989, 1, x_987); +return x_989; } } } } else { -lean_object* x_988; uint8_t x_989; -lean_dec(x_909); -x_988 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__36; -x_989 = lean_string_dec_eq(x_908, x_988); -lean_dec(x_908); -if (x_989 == 0) +lean_object* x_990; uint8_t x_991; +lean_dec(x_911); +x_990 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__36; +x_991 = lean_string_dec_eq(x_910, x_990); +lean_dec(x_910); +if (x_991 == 0) { -lean_object* x_990; lean_object* x_991; lean_object* x_992; -lean_dec(x_907); +lean_object* x_992; lean_object* x_993; lean_object* x_994; +lean_dec(x_909); +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_3); -x_990 = lean_unsigned_to_nat(0u); -x_991 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_991, 0, x_2); -lean_ctor_set(x_991, 1, x_990); -x_992 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_992, 0, x_991); -lean_ctor_set(x_992, 1, x_13); -return x_992; +x_992 = lean_unsigned_to_nat(0u); +x_993 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_993, 0, x_2); +lean_ctor_set(x_993, 1, x_992); +x_994 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_994, 0, x_993); +lean_ctor_set(x_994, 1, x_15); +return x_994; } else { -lean_object* x_993; lean_object* x_994; uint8_t x_995; -x_993 = lean_array_get_size(x_907); -x_994 = lean_unsigned_to_nat(4u); -x_995 = lean_nat_dec_eq(x_993, x_994); -lean_dec(x_993); -if (x_995 == 0) +lean_object* x_995; lean_object* x_996; uint8_t x_997; +x_995 = lean_array_get_size(x_909); +x_996 = lean_unsigned_to_nat(4u); +x_997 = lean_nat_dec_eq(x_995, x_996); +lean_dec(x_995); +if (x_997 == 0) { -lean_object* x_996; lean_object* x_997; lean_object* x_998; -lean_dec(x_907); +lean_object* x_998; lean_object* x_999; lean_object* x_1000; +lean_dec(x_909); +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_3); -x_996 = lean_unsigned_to_nat(0u); -x_997 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_997, 0, x_2); -lean_ctor_set(x_997, 1, x_996); -x_998 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_998, 0, x_997); -lean_ctor_set(x_998, 1, x_13); -return x_998; +x_998 = lean_unsigned_to_nat(0u); +x_999 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_999, 0, x_2); +lean_ctor_set(x_999, 1, x_998); +x_1000 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1000, 0, x_999); +lean_ctor_set(x_1000, 1, x_15); +return x_1000; } else { -lean_object* x_999; lean_object* x_1000; -x_999 = lean_unsigned_to_nat(0u); -x_1000 = lean_array_fget(x_907, x_999); -switch (lean_obj_tag(x_1000)) { +lean_object* x_1001; lean_object* x_1002; +x_1001 = lean_unsigned_to_nat(0u); +x_1002 = lean_array_fget(x_909, x_1001); +switch (lean_obj_tag(x_1002)) { case 4: { -lean_object* x_1001; -x_1001 = lean_ctor_get(x_1000, 0); -lean_inc(x_1001); -if (lean_obj_tag(x_1001) == 1) -{ -lean_object* x_1002; -x_1002 = lean_ctor_get(x_1001, 0); -lean_inc(x_1002); -if (lean_obj_tag(x_1002) == 0) -{ -lean_object* x_1003; lean_object* x_1004; lean_object* x_1005; uint8_t x_1006; -x_1003 = lean_ctor_get(x_1000, 1); +lean_object* x_1003; +x_1003 = lean_ctor_get(x_1002, 0); lean_inc(x_1003); -lean_dec(x_1000); -x_1004 = lean_ctor_get(x_1001, 1); +if (lean_obj_tag(x_1003) == 1) +{ +lean_object* x_1004; +x_1004 = lean_ctor_get(x_1003, 0); lean_inc(x_1004); -lean_dec(x_1001); -x_1005 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__9; -x_1006 = lean_string_dec_eq(x_1004, x_1005); -if (x_1006 == 0) -{ -lean_object* x_1007; uint8_t x_1008; -x_1007 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__1; -x_1008 = lean_string_dec_eq(x_1004, x_1007); -lean_dec(x_1004); -if (x_1008 == 0) +if (lean_obj_tag(x_1004) == 0) { -lean_object* x_1009; lean_object* x_1010; +lean_object* x_1005; lean_object* x_1006; lean_object* x_1007; uint8_t x_1008; +x_1005 = lean_ctor_get(x_1002, 1); +lean_inc(x_1005); +lean_dec(x_1002); +x_1006 = lean_ctor_get(x_1003, 1); +lean_inc(x_1006); lean_dec(x_1003); -lean_dec(x_907); +x_1007 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__9; +x_1008 = lean_string_dec_eq(x_1006, x_1007); +if (x_1008 == 0) +{ +lean_object* x_1009; uint8_t x_1010; +x_1009 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__1; +x_1010 = lean_string_dec_eq(x_1006, x_1009); +lean_dec(x_1006); +if (x_1010 == 0) +{ +lean_object* x_1011; lean_object* x_1012; +lean_dec(x_1005); +lean_dec(x_909); +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_3); -x_1009 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1009, 0, x_2); -lean_ctor_set(x_1009, 1, x_999); -x_1010 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1010, 0, x_1009); -lean_ctor_set(x_1010, 1, x_13); -return x_1010; +x_1011 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1011, 0, x_2); +lean_ctor_set(x_1011, 1, x_1001); +x_1012 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1012, 0, x_1011); +lean_ctor_set(x_1012, 1, x_15); +return x_1012; } else { -if (lean_obj_tag(x_1003) == 0) +if (lean_obj_tag(x_1005) == 0) { -lean_object* x_1011; lean_object* x_1012; lean_object* x_1013; lean_object* x_1014; lean_object* x_1015; lean_object* x_1016; lean_object* x_1017; -x_1011 = lean_unsigned_to_nat(2u); -x_1012 = lean_array_fget(x_907, x_1011); -x_1013 = lean_unsigned_to_nat(3u); -x_1014 = lean_array_fget(x_907, x_1013); -lean_dec(x_907); -x_1015 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__39; -x_1016 = l_Lean_mkApp3(x_1015, x_1012, x_1014, x_3); -x_1017 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_1016, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_1017; +lean_object* x_1013; lean_object* x_1014; lean_object* x_1015; lean_object* x_1016; lean_object* x_1017; lean_object* x_1018; lean_object* x_1019; +x_1013 = lean_unsigned_to_nat(2u); +x_1014 = lean_array_fget(x_909, x_1013); +x_1015 = lean_unsigned_to_nat(3u); +x_1016 = lean_array_fget(x_909, x_1015); +lean_dec(x_909); +x_1017 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__39; +x_1018 = l_Lean_mkApp3(x_1017, x_1014, x_1016, x_3); +x_1019 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_1018, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_1019; } else { -lean_object* x_1018; lean_object* x_1019; -lean_dec(x_1003); -lean_dec(x_907); +lean_object* x_1020; lean_object* x_1021; +lean_dec(x_1005); +lean_dec(x_909); +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_3); -x_1018 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1018, 0, x_2); -lean_ctor_set(x_1018, 1, x_999); -x_1019 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1019, 0, x_1018); -lean_ctor_set(x_1019, 1, x_13); -return x_1019; +x_1020 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1020, 0, x_2); +lean_ctor_set(x_1020, 1, x_1001); +x_1021 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1021, 0, x_1020); +lean_ctor_set(x_1021, 1, x_15); +return x_1021; } } } else { -lean_dec(x_1004); -if (lean_obj_tag(x_1003) == 0) -{ -lean_object* x_1020; lean_object* x_1021; lean_object* x_1022; lean_object* x_1023; lean_object* x_1024; lean_object* x_1025; lean_object* x_1026; -x_1020 = lean_unsigned_to_nat(2u); -x_1021 = lean_array_fget(x_907, x_1020); -x_1022 = lean_unsigned_to_nat(3u); -x_1023 = lean_array_fget(x_907, x_1022); -lean_dec(x_907); -x_1024 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__41; -x_1025 = l_Lean_mkApp3(x_1024, x_1021, x_1023, x_3); -x_1026 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_1025, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_1026; -} -else +lean_dec(x_1006); +if (lean_obj_tag(x_1005) == 0) { -lean_object* x_1027; lean_object* x_1028; -lean_dec(x_1003); -lean_dec(x_907); -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_3); -x_1027 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1027, 0, x_2); -lean_ctor_set(x_1027, 1, x_999); -x_1028 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1028, 0, x_1027); -lean_ctor_set(x_1028, 1, x_13); +lean_object* x_1022; lean_object* x_1023; lean_object* x_1024; lean_object* x_1025; lean_object* x_1026; lean_object* x_1027; lean_object* x_1028; +x_1022 = lean_unsigned_to_nat(2u); +x_1023 = lean_array_fget(x_909, x_1022); +x_1024 = lean_unsigned_to_nat(3u); +x_1025 = lean_array_fget(x_909, x_1024); +lean_dec(x_909); +x_1026 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__41; +x_1027 = l_Lean_mkApp3(x_1026, x_1023, x_1025, x_3); +x_1028 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_1027, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); return x_1028; } -} -} else { lean_object* x_1029; lean_object* x_1030; -lean_dec(x_1002); -lean_dec(x_1001); -lean_dec(x_1000); -lean_dec(x_907); +lean_dec(x_1005); +lean_dec(x_909); +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_3); x_1029 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1029, 0, x_2); -lean_ctor_set(x_1029, 1, x_999); +lean_ctor_set(x_1029, 1, x_1001); x_1030 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1030, 0, x_1029); -lean_ctor_set(x_1030, 1, x_13); +lean_ctor_set(x_1030, 1, x_15); return x_1030; } } +} else { lean_object* x_1031; lean_object* x_1032; -lean_dec(x_1001); -lean_dec(x_1000); -lean_dec(x_907); +lean_dec(x_1004); +lean_dec(x_1003); +lean_dec(x_1002); +lean_dec(x_909); +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_3); x_1031 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1031, 0, x_2); -lean_ctor_set(x_1031, 1, x_999); +lean_ctor_set(x_1031, 1, x_1001); x_1032 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1032, 0, x_1031); -lean_ctor_set(x_1032, 1, x_13); +lean_ctor_set(x_1032, 1, x_15); return x_1032; } } -case 5: -{ -lean_object* x_1033; -x_1033 = lean_ctor_get(x_1000, 0); -lean_inc(x_1033); -if (lean_obj_tag(x_1033) == 4) -{ -lean_object* x_1034; -x_1034 = lean_ctor_get(x_1033, 0); -lean_inc(x_1034); -if (lean_obj_tag(x_1034) == 1) -{ -lean_object* x_1035; -x_1035 = lean_ctor_get(x_1034, 0); -lean_inc(x_1035); -if (lean_obj_tag(x_1035) == 0) -{ -lean_object* x_1036; lean_object* x_1037; lean_object* x_1038; lean_object* x_1039; uint8_t x_1040; -x_1036 = lean_ctor_get(x_1000, 1); -lean_inc(x_1036); -lean_dec(x_1000); -x_1037 = lean_ctor_get(x_1033, 1); -lean_inc(x_1037); -lean_dec(x_1033); -x_1038 = lean_ctor_get(x_1034, 1); -lean_inc(x_1038); -lean_dec(x_1034); -x_1039 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__4; -x_1040 = lean_string_dec_eq(x_1038, x_1039); -lean_dec(x_1038); -if (x_1040 == 0) +else { -lean_object* x_1041; lean_object* x_1042; -lean_dec(x_1037); -lean_dec(x_1036); -lean_dec(x_907); +lean_object* x_1033; lean_object* x_1034; +lean_dec(x_1003); +lean_dec(x_1002); +lean_dec(x_909); +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_3); -x_1041 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1041, 0, x_2); -lean_ctor_set(x_1041, 1, x_999); -x_1042 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1042, 0, x_1041); -lean_ctor_set(x_1042, 1, x_13); -return x_1042; +x_1033 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1033, 0, x_2); +lean_ctor_set(x_1033, 1, x_1001); +x_1034 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1034, 0, x_1033); +lean_ctor_set(x_1034, 1, x_15); +return x_1034; } -else +} +case 5: { -if (lean_obj_tag(x_1037) == 0) +lean_object* x_1035; +x_1035 = lean_ctor_get(x_1002, 0); +lean_inc(x_1035); +if (lean_obj_tag(x_1035) == 4) { -lean_object* x_1043; lean_object* x_1044; lean_object* x_1045; lean_object* x_1046; lean_object* x_1047; lean_object* x_1048; lean_object* x_1049; -x_1043 = lean_unsigned_to_nat(2u); -x_1044 = lean_array_fget(x_907, x_1043); -x_1045 = lean_unsigned_to_nat(3u); -x_1046 = lean_array_fget(x_907, x_1045); -lean_dec(x_907); -x_1047 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__44; -x_1048 = l_Lean_mkApp4(x_1047, x_1036, x_1044, x_1046, x_3); -x_1049 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_1048, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_1049; -} -else +lean_object* x_1036; +x_1036 = lean_ctor_get(x_1035, 0); +lean_inc(x_1036); +if (lean_obj_tag(x_1036) == 1) { -lean_object* x_1050; lean_object* x_1051; -lean_dec(x_1037); +lean_object* x_1037; +x_1037 = lean_ctor_get(x_1036, 0); +lean_inc(x_1037); +if (lean_obj_tag(x_1037) == 0) +{ +lean_object* x_1038; lean_object* x_1039; lean_object* x_1040; lean_object* x_1041; uint8_t x_1042; +x_1038 = lean_ctor_get(x_1002, 1); +lean_inc(x_1038); +lean_dec(x_1002); +x_1039 = lean_ctor_get(x_1035, 1); +lean_inc(x_1039); +lean_dec(x_1035); +x_1040 = lean_ctor_get(x_1036, 1); +lean_inc(x_1040); lean_dec(x_1036); -lean_dec(x_907); +x_1041 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__4; +x_1042 = lean_string_dec_eq(x_1040, x_1041); +lean_dec(x_1040); +if (x_1042 == 0) +{ +lean_object* x_1043; lean_object* x_1044; +lean_dec(x_1039); +lean_dec(x_1038); +lean_dec(x_909); +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_3); -x_1050 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1050, 0, x_2); -lean_ctor_set(x_1050, 1, x_999); -x_1051 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1051, 0, x_1050); -lean_ctor_set(x_1051, 1, x_13); -return x_1051; -} +x_1043 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1043, 0, x_2); +lean_ctor_set(x_1043, 1, x_1001); +x_1044 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1044, 0, x_1043); +lean_ctor_set(x_1044, 1, x_15); +return x_1044; } +else +{ +if (lean_obj_tag(x_1039) == 0) +{ +lean_object* x_1045; lean_object* x_1046; lean_object* x_1047; lean_object* x_1048; lean_object* x_1049; lean_object* x_1050; lean_object* x_1051; +x_1045 = lean_unsigned_to_nat(2u); +x_1046 = lean_array_fget(x_909, x_1045); +x_1047 = lean_unsigned_to_nat(3u); +x_1048 = lean_array_fget(x_909, x_1047); +lean_dec(x_909); +x_1049 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__44; +x_1050 = l_Lean_mkApp4(x_1049, x_1038, x_1046, x_1048, x_3); +x_1051 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_1050, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_1051; } else { lean_object* x_1052; lean_object* x_1053; -lean_dec(x_1035); -lean_dec(x_1034); -lean_dec(x_1033); -lean_dec(x_1000); -lean_dec(x_907); +lean_dec(x_1039); +lean_dec(x_1038); +lean_dec(x_909); +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_3); x_1052 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1052, 0, x_2); -lean_ctor_set(x_1052, 1, x_999); +lean_ctor_set(x_1052, 1, x_1001); x_1053 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1053, 0, x_1052); -lean_ctor_set(x_1053, 1, x_13); +lean_ctor_set(x_1053, 1, x_15); return x_1053; } } +} else { lean_object* x_1054; lean_object* x_1055; -lean_dec(x_1034); -lean_dec(x_1033); -lean_dec(x_1000); -lean_dec(x_907); +lean_dec(x_1037); +lean_dec(x_1036); +lean_dec(x_1035); +lean_dec(x_1002); +lean_dec(x_909); +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_3); x_1054 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1054, 0, x_2); -lean_ctor_set(x_1054, 1, x_999); +lean_ctor_set(x_1054, 1, x_1001); x_1055 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1055, 0, x_1054); -lean_ctor_set(x_1055, 1, x_13); +lean_ctor_set(x_1055, 1, x_15); return x_1055; } } else { lean_object* x_1056; lean_object* x_1057; -lean_dec(x_1033); -lean_dec(x_1000); -lean_dec(x_907); +lean_dec(x_1036); +lean_dec(x_1035); +lean_dec(x_1002); +lean_dec(x_909); +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_3); x_1056 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1056, 0, x_2); -lean_ctor_set(x_1056, 1, x_999); +lean_ctor_set(x_1056, 1, x_1001); x_1057 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1057, 0, x_1056); -lean_ctor_set(x_1057, 1, x_13); +lean_ctor_set(x_1057, 1, x_15); return x_1057; } } -default: +else { lean_object* x_1058; lean_object* x_1059; -lean_dec(x_1000); -lean_dec(x_907); +lean_dec(x_1035); +lean_dec(x_1002); +lean_dec(x_909); +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_3); x_1058 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1058, 0, x_2); -lean_ctor_set(x_1058, 1, x_999); +lean_ctor_set(x_1058, 1, x_1001); x_1059 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1059, 0, x_1058); -lean_ctor_set(x_1059, 1, x_13); +lean_ctor_set(x_1059, 1, x_15); return x_1059; } } +default: +{ +lean_object* x_1060; lean_object* x_1061; +lean_dec(x_1002); +lean_dec(x_909); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_3); +x_1060 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1060, 0, x_2); +lean_ctor_set(x_1060, 1, x_1001); +x_1061 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1061, 0, x_1060); +lean_ctor_set(x_1061, 1, x_15); +return x_1061; +} +} } } } } else { -lean_object* x_1060; uint8_t x_1061; -lean_dec(x_909); -x_1060 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__45; -x_1061 = lean_string_dec_eq(x_908, x_1060); -lean_dec(x_908); -if (x_1061 == 0) +lean_object* x_1062; uint8_t x_1063; +lean_dec(x_911); +x_1062 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__45; +x_1063 = lean_string_dec_eq(x_910, x_1062); +lean_dec(x_910); +if (x_1063 == 0) { -lean_object* x_1062; lean_object* x_1063; lean_object* x_1064; -lean_dec(x_907); +lean_object* x_1064; lean_object* x_1065; lean_object* x_1066; +lean_dec(x_909); +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_3); -x_1062 = lean_unsigned_to_nat(0u); -x_1063 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1063, 0, x_2); -lean_ctor_set(x_1063, 1, x_1062); -x_1064 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1064, 0, x_1063); -lean_ctor_set(x_1064, 1, x_13); -return x_1064; +x_1064 = lean_unsigned_to_nat(0u); +x_1065 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1065, 0, x_2); +lean_ctor_set(x_1065, 1, x_1064); +x_1066 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1066, 0, x_1065); +lean_ctor_set(x_1066, 1, x_15); +return x_1066; } else { -lean_object* x_1065; lean_object* x_1066; uint8_t x_1067; -x_1065 = lean_array_get_size(x_907); -x_1066 = lean_unsigned_to_nat(4u); -x_1067 = lean_nat_dec_eq(x_1065, x_1066); -lean_dec(x_1065); -if (x_1067 == 0) +lean_object* x_1067; lean_object* x_1068; uint8_t x_1069; +x_1067 = lean_array_get_size(x_909); +x_1068 = lean_unsigned_to_nat(4u); +x_1069 = lean_nat_dec_eq(x_1067, x_1068); +lean_dec(x_1067); +if (x_1069 == 0) { -lean_object* x_1068; lean_object* x_1069; lean_object* x_1070; -lean_dec(x_907); +lean_object* x_1070; lean_object* x_1071; lean_object* x_1072; +lean_dec(x_909); +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_3); -x_1068 = lean_unsigned_to_nat(0u); -x_1069 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1069, 0, x_2); -lean_ctor_set(x_1069, 1, x_1068); -x_1070 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1070, 0, x_1069); -lean_ctor_set(x_1070, 1, x_13); -return x_1070; +x_1070 = lean_unsigned_to_nat(0u); +x_1071 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1071, 0, x_2); +lean_ctor_set(x_1071, 1, x_1070); +x_1072 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1072, 0, x_1071); +lean_ctor_set(x_1072, 1, x_15); +return x_1072; } else { -lean_object* x_1071; lean_object* x_1072; -x_1071 = lean_unsigned_to_nat(0u); -x_1072 = lean_array_fget(x_907, x_1071); -switch (lean_obj_tag(x_1072)) { +lean_object* x_1073; lean_object* x_1074; +x_1073 = lean_unsigned_to_nat(0u); +x_1074 = lean_array_fget(x_909, x_1073); +switch (lean_obj_tag(x_1074)) { case 4: { -lean_object* x_1073; -x_1073 = lean_ctor_get(x_1072, 0); -lean_inc(x_1073); -if (lean_obj_tag(x_1073) == 1) -{ -lean_object* x_1074; -x_1074 = lean_ctor_get(x_1073, 0); -lean_inc(x_1074); -if (lean_obj_tag(x_1074) == 0) -{ -lean_object* x_1075; lean_object* x_1076; lean_object* x_1077; uint8_t x_1078; -x_1075 = lean_ctor_get(x_1072, 1); +lean_object* x_1075; +x_1075 = lean_ctor_get(x_1074, 0); lean_inc(x_1075); -lean_dec(x_1072); -x_1076 = lean_ctor_get(x_1073, 1); +if (lean_obj_tag(x_1075) == 1) +{ +lean_object* x_1076; +x_1076 = lean_ctor_get(x_1075, 0); lean_inc(x_1076); -lean_dec(x_1073); -x_1077 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__9; -x_1078 = lean_string_dec_eq(x_1076, x_1077); -if (x_1078 == 0) -{ -lean_object* x_1079; uint8_t x_1080; -x_1079 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__1; -x_1080 = lean_string_dec_eq(x_1076, x_1079); -lean_dec(x_1076); -if (x_1080 == 0) +if (lean_obj_tag(x_1076) == 0) { -lean_object* x_1081; lean_object* x_1082; +lean_object* x_1077; lean_object* x_1078; lean_object* x_1079; uint8_t x_1080; +x_1077 = lean_ctor_get(x_1074, 1); +lean_inc(x_1077); +lean_dec(x_1074); +x_1078 = lean_ctor_get(x_1075, 1); +lean_inc(x_1078); lean_dec(x_1075); -lean_dec(x_907); +x_1079 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__9; +x_1080 = lean_string_dec_eq(x_1078, x_1079); +if (x_1080 == 0) +{ +lean_object* x_1081; uint8_t x_1082; +x_1081 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__1; +x_1082 = lean_string_dec_eq(x_1078, x_1081); +lean_dec(x_1078); +if (x_1082 == 0) +{ +lean_object* x_1083; lean_object* x_1084; +lean_dec(x_1077); +lean_dec(x_909); +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_3); -x_1081 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1081, 0, x_2); -lean_ctor_set(x_1081, 1, x_1071); -x_1082 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1082, 0, x_1081); -lean_ctor_set(x_1082, 1, x_13); -return x_1082; +x_1083 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1083, 0, x_2); +lean_ctor_set(x_1083, 1, x_1073); +x_1084 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1084, 0, x_1083); +lean_ctor_set(x_1084, 1, x_15); +return x_1084; } else { -if (lean_obj_tag(x_1075) == 0) +if (lean_obj_tag(x_1077) == 0) { -lean_object* x_1083; lean_object* x_1084; lean_object* x_1085; lean_object* x_1086; lean_object* x_1087; lean_object* x_1088; lean_object* x_1089; -x_1083 = lean_unsigned_to_nat(2u); -x_1084 = lean_array_fget(x_907, x_1083); -x_1085 = lean_unsigned_to_nat(3u); -x_1086 = lean_array_fget(x_907, x_1085); -lean_dec(x_907); -x_1087 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__48; -x_1088 = l_Lean_mkApp3(x_1087, x_1084, x_1086, x_3); -x_1089 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_1088, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_1089; +lean_object* x_1085; lean_object* x_1086; lean_object* x_1087; lean_object* x_1088; lean_object* x_1089; lean_object* x_1090; lean_object* x_1091; +x_1085 = lean_unsigned_to_nat(2u); +x_1086 = lean_array_fget(x_909, x_1085); +x_1087 = lean_unsigned_to_nat(3u); +x_1088 = lean_array_fget(x_909, x_1087); +lean_dec(x_909); +x_1089 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__48; +x_1090 = l_Lean_mkApp3(x_1089, x_1086, x_1088, x_3); +x_1091 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_1090, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_1091; } else { -lean_object* x_1090; lean_object* x_1091; -lean_dec(x_1075); -lean_dec(x_907); +lean_object* x_1092; lean_object* x_1093; +lean_dec(x_1077); +lean_dec(x_909); +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_3); -x_1090 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1090, 0, x_2); -lean_ctor_set(x_1090, 1, x_1071); -x_1091 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1091, 0, x_1090); -lean_ctor_set(x_1091, 1, x_13); -return x_1091; +x_1092 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1092, 0, x_2); +lean_ctor_set(x_1092, 1, x_1073); +x_1093 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1093, 0, x_1092); +lean_ctor_set(x_1093, 1, x_15); +return x_1093; } } } else { -lean_dec(x_1076); -if (lean_obj_tag(x_1075) == 0) -{ -lean_object* x_1092; lean_object* x_1093; lean_object* x_1094; lean_object* x_1095; lean_object* x_1096; lean_object* x_1097; lean_object* x_1098; -x_1092 = lean_unsigned_to_nat(2u); -x_1093 = lean_array_fget(x_907, x_1092); -x_1094 = lean_unsigned_to_nat(3u); -x_1095 = lean_array_fget(x_907, x_1094); -lean_dec(x_907); -x_1096 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__50; -x_1097 = l_Lean_mkApp3(x_1096, x_1093, x_1095, x_3); -x_1098 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_1097, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_1098; -} -else +lean_dec(x_1078); +if (lean_obj_tag(x_1077) == 0) { -lean_object* x_1099; lean_object* x_1100; -lean_dec(x_1075); -lean_dec(x_907); -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_3); -x_1099 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1099, 0, x_2); -lean_ctor_set(x_1099, 1, x_1071); -x_1100 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1100, 0, x_1099); -lean_ctor_set(x_1100, 1, x_13); +lean_object* x_1094; lean_object* x_1095; lean_object* x_1096; lean_object* x_1097; lean_object* x_1098; lean_object* x_1099; lean_object* x_1100; +x_1094 = lean_unsigned_to_nat(2u); +x_1095 = lean_array_fget(x_909, x_1094); +x_1096 = lean_unsigned_to_nat(3u); +x_1097 = lean_array_fget(x_909, x_1096); +lean_dec(x_909); +x_1098 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__50; +x_1099 = l_Lean_mkApp3(x_1098, x_1095, x_1097, x_3); +x_1100 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_1099, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); return x_1100; } -} -} else { lean_object* x_1101; lean_object* x_1102; -lean_dec(x_1074); -lean_dec(x_1073); -lean_dec(x_1072); -lean_dec(x_907); +lean_dec(x_1077); +lean_dec(x_909); +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_3); x_1101 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1101, 0, x_2); -lean_ctor_set(x_1101, 1, x_1071); +lean_ctor_set(x_1101, 1, x_1073); x_1102 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1102, 0, x_1101); -lean_ctor_set(x_1102, 1, x_13); +lean_ctor_set(x_1102, 1, x_15); return x_1102; } } +} else { lean_object* x_1103; lean_object* x_1104; -lean_dec(x_1073); -lean_dec(x_1072); -lean_dec(x_907); +lean_dec(x_1076); +lean_dec(x_1075); +lean_dec(x_1074); +lean_dec(x_909); +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_3); x_1103 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1103, 0, x_2); -lean_ctor_set(x_1103, 1, x_1071); +lean_ctor_set(x_1103, 1, x_1073); x_1104 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1104, 0, x_1103); -lean_ctor_set(x_1104, 1, x_13); +lean_ctor_set(x_1104, 1, x_15); return x_1104; } } -case 5: -{ -lean_object* x_1105; -x_1105 = lean_ctor_get(x_1072, 0); -lean_inc(x_1105); -if (lean_obj_tag(x_1105) == 4) -{ -lean_object* x_1106; -x_1106 = lean_ctor_get(x_1105, 0); -lean_inc(x_1106); -if (lean_obj_tag(x_1106) == 1) -{ -lean_object* x_1107; -x_1107 = lean_ctor_get(x_1106, 0); -lean_inc(x_1107); -if (lean_obj_tag(x_1107) == 0) -{ -lean_object* x_1108; lean_object* x_1109; lean_object* x_1110; lean_object* x_1111; uint8_t x_1112; -x_1108 = lean_ctor_get(x_1072, 1); -lean_inc(x_1108); -lean_dec(x_1072); -x_1109 = lean_ctor_get(x_1105, 1); -lean_inc(x_1109); -lean_dec(x_1105); -x_1110 = lean_ctor_get(x_1106, 1); -lean_inc(x_1110); -lean_dec(x_1106); -x_1111 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__4; -x_1112 = lean_string_dec_eq(x_1110, x_1111); -lean_dec(x_1110); -if (x_1112 == 0) +else { -lean_object* x_1113; lean_object* x_1114; -lean_dec(x_1109); -lean_dec(x_1108); -lean_dec(x_907); +lean_object* x_1105; lean_object* x_1106; +lean_dec(x_1075); +lean_dec(x_1074); +lean_dec(x_909); +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_3); -x_1113 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1113, 0, x_2); -lean_ctor_set(x_1113, 1, x_1071); -x_1114 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1114, 0, x_1113); -lean_ctor_set(x_1114, 1, x_13); -return x_1114; +x_1105 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1105, 0, x_2); +lean_ctor_set(x_1105, 1, x_1073); +x_1106 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1106, 0, x_1105); +lean_ctor_set(x_1106, 1, x_15); +return x_1106; } -else +} +case 5: { -if (lean_obj_tag(x_1109) == 0) +lean_object* x_1107; +x_1107 = lean_ctor_get(x_1074, 0); +lean_inc(x_1107); +if (lean_obj_tag(x_1107) == 4) { -lean_object* x_1115; lean_object* x_1116; lean_object* x_1117; lean_object* x_1118; lean_object* x_1119; lean_object* x_1120; lean_object* x_1121; -x_1115 = lean_unsigned_to_nat(2u); -x_1116 = lean_array_fget(x_907, x_1115); -x_1117 = lean_unsigned_to_nat(3u); -x_1118 = lean_array_fget(x_907, x_1117); -lean_dec(x_907); -x_1119 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__53; -x_1120 = l_Lean_mkApp4(x_1119, x_1108, x_1116, x_1118, x_3); -x_1121 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_1120, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_1121; -} -else +lean_object* x_1108; +x_1108 = lean_ctor_get(x_1107, 0); +lean_inc(x_1108); +if (lean_obj_tag(x_1108) == 1) { -lean_object* x_1122; lean_object* x_1123; -lean_dec(x_1109); +lean_object* x_1109; +x_1109 = lean_ctor_get(x_1108, 0); +lean_inc(x_1109); +if (lean_obj_tag(x_1109) == 0) +{ +lean_object* x_1110; lean_object* x_1111; lean_object* x_1112; lean_object* x_1113; uint8_t x_1114; +x_1110 = lean_ctor_get(x_1074, 1); +lean_inc(x_1110); +lean_dec(x_1074); +x_1111 = lean_ctor_get(x_1107, 1); +lean_inc(x_1111); +lean_dec(x_1107); +x_1112 = lean_ctor_get(x_1108, 1); +lean_inc(x_1112); lean_dec(x_1108); -lean_dec(x_907); +x_1113 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__4; +x_1114 = lean_string_dec_eq(x_1112, x_1113); +lean_dec(x_1112); +if (x_1114 == 0) +{ +lean_object* x_1115; lean_object* x_1116; +lean_dec(x_1111); +lean_dec(x_1110); +lean_dec(x_909); +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_3); -x_1122 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1122, 0, x_2); -lean_ctor_set(x_1122, 1, x_1071); -x_1123 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1123, 0, x_1122); -lean_ctor_set(x_1123, 1, x_13); -return x_1123; -} +x_1115 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1115, 0, x_2); +lean_ctor_set(x_1115, 1, x_1073); +x_1116 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1116, 0, x_1115); +lean_ctor_set(x_1116, 1, x_15); +return x_1116; } +else +{ +if (lean_obj_tag(x_1111) == 0) +{ +lean_object* x_1117; lean_object* x_1118; lean_object* x_1119; lean_object* x_1120; lean_object* x_1121; lean_object* x_1122; lean_object* x_1123; +x_1117 = lean_unsigned_to_nat(2u); +x_1118 = lean_array_fget(x_909, x_1117); +x_1119 = lean_unsigned_to_nat(3u); +x_1120 = lean_array_fget(x_909, x_1119); +lean_dec(x_909); +x_1121 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__53; +x_1122 = l_Lean_mkApp4(x_1121, x_1110, x_1118, x_1120, x_3); +x_1123 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_1122, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_1123; } else { lean_object* x_1124; lean_object* x_1125; -lean_dec(x_1107); -lean_dec(x_1106); -lean_dec(x_1105); -lean_dec(x_1072); -lean_dec(x_907); +lean_dec(x_1111); +lean_dec(x_1110); +lean_dec(x_909); +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_3); x_1124 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1124, 0, x_2); -lean_ctor_set(x_1124, 1, x_1071); +lean_ctor_set(x_1124, 1, x_1073); x_1125 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1125, 0, x_1124); -lean_ctor_set(x_1125, 1, x_13); +lean_ctor_set(x_1125, 1, x_15); return x_1125; } } +} else { lean_object* x_1126; lean_object* x_1127; -lean_dec(x_1106); -lean_dec(x_1105); -lean_dec(x_1072); -lean_dec(x_907); +lean_dec(x_1109); +lean_dec(x_1108); +lean_dec(x_1107); +lean_dec(x_1074); +lean_dec(x_909); +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_3); x_1126 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1126, 0, x_2); -lean_ctor_set(x_1126, 1, x_1071); +lean_ctor_set(x_1126, 1, x_1073); x_1127 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1127, 0, x_1126); -lean_ctor_set(x_1127, 1, x_13); +lean_ctor_set(x_1127, 1, x_15); return x_1127; } } else { lean_object* x_1128; lean_object* x_1129; -lean_dec(x_1105); -lean_dec(x_1072); -lean_dec(x_907); +lean_dec(x_1108); +lean_dec(x_1107); +lean_dec(x_1074); +lean_dec(x_909); +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_3); x_1128 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1128, 0, x_2); -lean_ctor_set(x_1128, 1, x_1071); +lean_ctor_set(x_1128, 1, x_1073); x_1129 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1129, 0, x_1128); -lean_ctor_set(x_1129, 1, x_13); +lean_ctor_set(x_1129, 1, x_15); return x_1129; } } -default: +else { lean_object* x_1130; lean_object* x_1131; -lean_dec(x_1072); -lean_dec(x_907); +lean_dec(x_1107); +lean_dec(x_1074); +lean_dec(x_909); +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_3); x_1130 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1130, 0, x_2); -lean_ctor_set(x_1130, 1, x_1071); +lean_ctor_set(x_1130, 1, x_1073); x_1131 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1131, 0, x_1130); -lean_ctor_set(x_1131, 1, x_13); +lean_ctor_set(x_1131, 1, x_15); return x_1131; } } +default: +{ +lean_object* x_1132; lean_object* x_1133; +lean_dec(x_1074); +lean_dec(x_909); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_3); +x_1132 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1132, 0, x_2); +lean_ctor_set(x_1132, 1, x_1073); +x_1133 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1133, 0, x_1132); +lean_ctor_set(x_1133, 1, x_15); +return x_1133; +} +} } } } } else { -lean_object* x_1132; uint8_t x_1133; -lean_dec(x_909); -x_1132 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__14; -x_1133 = lean_string_dec_eq(x_908, x_1132); -lean_dec(x_908); -if (x_1133 == 0) +lean_object* x_1134; uint8_t x_1135; +lean_dec(x_911); +x_1134 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__14; +x_1135 = lean_string_dec_eq(x_910, x_1134); +lean_dec(x_910); +if (x_1135 == 0) { -lean_object* x_1134; lean_object* x_1135; lean_object* x_1136; -lean_dec(x_907); +lean_object* x_1136; lean_object* x_1137; lean_object* x_1138; +lean_dec(x_909); +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_3); -x_1134 = lean_unsigned_to_nat(0u); -x_1135 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1135, 0, x_2); -lean_ctor_set(x_1135, 1, x_1134); -x_1136 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1136, 0, x_1135); -lean_ctor_set(x_1136, 1, x_13); -return x_1136; +x_1136 = lean_unsigned_to_nat(0u); +x_1137 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1137, 0, x_2); +lean_ctor_set(x_1137, 1, x_1136); +x_1138 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1138, 0, x_1137); +lean_ctor_set(x_1138, 1, x_15); +return x_1138; } else { -lean_object* x_1137; lean_object* x_1138; uint8_t x_1139; -x_1137 = lean_array_get_size(x_907); -x_1138 = lean_unsigned_to_nat(4u); -x_1139 = lean_nat_dec_eq(x_1137, x_1138); -lean_dec(x_1137); -if (x_1139 == 0) +lean_object* x_1139; lean_object* x_1140; uint8_t x_1141; +x_1139 = lean_array_get_size(x_909); +x_1140 = lean_unsigned_to_nat(4u); +x_1141 = lean_nat_dec_eq(x_1139, x_1140); +lean_dec(x_1139); +if (x_1141 == 0) { -lean_object* x_1140; lean_object* x_1141; lean_object* x_1142; -lean_dec(x_907); +lean_object* x_1142; lean_object* x_1143; lean_object* x_1144; +lean_dec(x_909); +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_3); -x_1140 = lean_unsigned_to_nat(0u); -x_1141 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1141, 0, x_2); -lean_ctor_set(x_1141, 1, x_1140); -x_1142 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1142, 0, x_1141); -lean_ctor_set(x_1142, 1, x_13); -return x_1142; +x_1142 = lean_unsigned_to_nat(0u); +x_1143 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1143, 0, x_2); +lean_ctor_set(x_1143, 1, x_1142); +x_1144 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1144, 0, x_1143); +lean_ctor_set(x_1144, 1, x_15); +return x_1144; } else { -lean_object* x_1143; lean_object* x_1144; -x_1143 = lean_unsigned_to_nat(0u); -x_1144 = lean_array_fget(x_907, x_1143); -switch (lean_obj_tag(x_1144)) { +lean_object* x_1145; lean_object* x_1146; +x_1145 = lean_unsigned_to_nat(0u); +x_1146 = lean_array_fget(x_909, x_1145); +switch (lean_obj_tag(x_1146)) { case 4: { -lean_object* x_1145; -x_1145 = lean_ctor_get(x_1144, 0); -lean_inc(x_1145); -if (lean_obj_tag(x_1145) == 1) -{ -lean_object* x_1146; -x_1146 = lean_ctor_get(x_1145, 0); -lean_inc(x_1146); -if (lean_obj_tag(x_1146) == 0) -{ -lean_object* x_1147; lean_object* x_1148; lean_object* x_1149; uint8_t x_1150; -x_1147 = lean_ctor_get(x_1144, 1); +lean_object* x_1147; +x_1147 = lean_ctor_get(x_1146, 0); lean_inc(x_1147); -lean_dec(x_1144); -x_1148 = lean_ctor_get(x_1145, 1); +if (lean_obj_tag(x_1147) == 1) +{ +lean_object* x_1148; +x_1148 = lean_ctor_get(x_1147, 0); lean_inc(x_1148); -lean_dec(x_1145); -x_1149 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__9; -x_1150 = lean_string_dec_eq(x_1148, x_1149); -if (x_1150 == 0) -{ -lean_object* x_1151; uint8_t x_1152; -x_1151 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__1; -x_1152 = lean_string_dec_eq(x_1148, x_1151); -lean_dec(x_1148); -if (x_1152 == 0) +if (lean_obj_tag(x_1148) == 0) { -lean_object* x_1153; lean_object* x_1154; +lean_object* x_1149; lean_object* x_1150; lean_object* x_1151; uint8_t x_1152; +x_1149 = lean_ctor_get(x_1146, 1); +lean_inc(x_1149); +lean_dec(x_1146); +x_1150 = lean_ctor_get(x_1147, 1); +lean_inc(x_1150); lean_dec(x_1147); -lean_dec(x_907); +x_1151 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__9; +x_1152 = lean_string_dec_eq(x_1150, x_1151); +if (x_1152 == 0) +{ +lean_object* x_1153; uint8_t x_1154; +x_1153 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__1; +x_1154 = lean_string_dec_eq(x_1150, x_1153); +lean_dec(x_1150); +if (x_1154 == 0) +{ +lean_object* x_1155; lean_object* x_1156; +lean_dec(x_1149); +lean_dec(x_909); +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_3); -x_1153 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1153, 0, x_2); -lean_ctor_set(x_1153, 1, x_1143); -x_1154 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1154, 0, x_1153); -lean_ctor_set(x_1154, 1, x_13); -return x_1154; +x_1155 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1155, 0, x_2); +lean_ctor_set(x_1155, 1, x_1145); +x_1156 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1156, 0, x_1155); +lean_ctor_set(x_1156, 1, x_15); +return x_1156; } else { -if (lean_obj_tag(x_1147) == 0) +if (lean_obj_tag(x_1149) == 0) { -lean_object* x_1155; lean_object* x_1156; lean_object* x_1157; lean_object* x_1158; lean_object* x_1159; lean_object* x_1160; lean_object* x_1161; -x_1155 = lean_unsigned_to_nat(2u); -x_1156 = lean_array_fget(x_907, x_1155); -x_1157 = lean_unsigned_to_nat(3u); -x_1158 = lean_array_fget(x_907, x_1157); -lean_dec(x_907); -x_1159 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__56; -x_1160 = l_Lean_mkApp3(x_1159, x_1156, x_1158, x_3); -x_1161 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_1160, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_1161; +lean_object* x_1157; lean_object* x_1158; lean_object* x_1159; lean_object* x_1160; lean_object* x_1161; lean_object* x_1162; lean_object* x_1163; +x_1157 = lean_unsigned_to_nat(2u); +x_1158 = lean_array_fget(x_909, x_1157); +x_1159 = lean_unsigned_to_nat(3u); +x_1160 = lean_array_fget(x_909, x_1159); +lean_dec(x_909); +x_1161 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__56; +x_1162 = l_Lean_mkApp3(x_1161, x_1158, x_1160, x_3); +x_1163 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_1162, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_1163; } else { -lean_object* x_1162; lean_object* x_1163; -lean_dec(x_1147); -lean_dec(x_907); +lean_object* x_1164; lean_object* x_1165; +lean_dec(x_1149); +lean_dec(x_909); +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_3); -x_1162 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1162, 0, x_2); -lean_ctor_set(x_1162, 1, x_1143); -x_1163 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1163, 0, x_1162); -lean_ctor_set(x_1163, 1, x_13); -return x_1163; +x_1164 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1164, 0, x_2); +lean_ctor_set(x_1164, 1, x_1145); +x_1165 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1165, 0, x_1164); +lean_ctor_set(x_1165, 1, x_15); +return x_1165; } } } else { -lean_dec(x_1148); -if (lean_obj_tag(x_1147) == 0) -{ -lean_object* x_1164; lean_object* x_1165; lean_object* x_1166; lean_object* x_1167; lean_object* x_1168; lean_object* x_1169; lean_object* x_1170; -x_1164 = lean_unsigned_to_nat(2u); -x_1165 = lean_array_fget(x_907, x_1164); -x_1166 = lean_unsigned_to_nat(3u); -x_1167 = lean_array_fget(x_907, x_1166); -lean_dec(x_907); -x_1168 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__59; -x_1169 = l_Lean_mkApp3(x_1168, x_1165, x_1167, x_3); -x_1170 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_1169, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_1170; -} -else +lean_dec(x_1150); +if (lean_obj_tag(x_1149) == 0) { -lean_object* x_1171; lean_object* x_1172; -lean_dec(x_1147); -lean_dec(x_907); -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_3); -x_1171 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1171, 0, x_2); -lean_ctor_set(x_1171, 1, x_1143); -x_1172 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1172, 0, x_1171); -lean_ctor_set(x_1172, 1, x_13); +lean_object* x_1166; lean_object* x_1167; lean_object* x_1168; lean_object* x_1169; lean_object* x_1170; lean_object* x_1171; lean_object* x_1172; +x_1166 = lean_unsigned_to_nat(2u); +x_1167 = lean_array_fget(x_909, x_1166); +x_1168 = lean_unsigned_to_nat(3u); +x_1169 = lean_array_fget(x_909, x_1168); +lean_dec(x_909); +x_1170 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__59; +x_1171 = l_Lean_mkApp3(x_1170, x_1167, x_1169, x_3); +x_1172 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_1171, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); return x_1172; } -} -} else { lean_object* x_1173; lean_object* x_1174; -lean_dec(x_1146); -lean_dec(x_1145); -lean_dec(x_1144); -lean_dec(x_907); +lean_dec(x_1149); +lean_dec(x_909); +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_3); x_1173 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1173, 0, x_2); -lean_ctor_set(x_1173, 1, x_1143); +lean_ctor_set(x_1173, 1, x_1145); x_1174 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1174, 0, x_1173); -lean_ctor_set(x_1174, 1, x_13); +lean_ctor_set(x_1174, 1, x_15); return x_1174; } } +} else { lean_object* x_1175; lean_object* x_1176; -lean_dec(x_1145); -lean_dec(x_1144); -lean_dec(x_907); +lean_dec(x_1148); +lean_dec(x_1147); +lean_dec(x_1146); +lean_dec(x_909); +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_3); x_1175 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1175, 0, x_2); -lean_ctor_set(x_1175, 1, x_1143); +lean_ctor_set(x_1175, 1, x_1145); x_1176 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1176, 0, x_1175); -lean_ctor_set(x_1176, 1, x_13); +lean_ctor_set(x_1176, 1, x_15); return x_1176; } } -case 5: -{ -lean_object* x_1177; -x_1177 = lean_ctor_get(x_1144, 0); -lean_inc(x_1177); -if (lean_obj_tag(x_1177) == 4) -{ -lean_object* x_1178; -x_1178 = lean_ctor_get(x_1177, 0); -lean_inc(x_1178); -if (lean_obj_tag(x_1178) == 1) -{ -lean_object* x_1179; -x_1179 = lean_ctor_get(x_1178, 0); -lean_inc(x_1179); -if (lean_obj_tag(x_1179) == 0) -{ -lean_object* x_1180; lean_object* x_1181; lean_object* x_1182; lean_object* x_1183; uint8_t x_1184; -x_1180 = lean_ctor_get(x_1144, 1); -lean_inc(x_1180); -lean_dec(x_1144); -x_1181 = lean_ctor_get(x_1177, 1); -lean_inc(x_1181); -lean_dec(x_1177); -x_1182 = lean_ctor_get(x_1178, 1); -lean_inc(x_1182); -lean_dec(x_1178); -x_1183 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__4; -x_1184 = lean_string_dec_eq(x_1182, x_1183); -lean_dec(x_1182); -if (x_1184 == 0) +else { -lean_object* x_1185; lean_object* x_1186; -lean_dec(x_1181); -lean_dec(x_1180); -lean_dec(x_907); +lean_object* x_1177; lean_object* x_1178; +lean_dec(x_1147); +lean_dec(x_1146); +lean_dec(x_909); +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_3); -x_1185 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1185, 0, x_2); -lean_ctor_set(x_1185, 1, x_1143); -x_1186 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1186, 0, x_1185); -lean_ctor_set(x_1186, 1, x_13); -return x_1186; +x_1177 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1177, 0, x_2); +lean_ctor_set(x_1177, 1, x_1145); +x_1178 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1178, 0, x_1177); +lean_ctor_set(x_1178, 1, x_15); +return x_1178; } -else +} +case 5: { -if (lean_obj_tag(x_1181) == 0) +lean_object* x_1179; +x_1179 = lean_ctor_get(x_1146, 0); +lean_inc(x_1179); +if (lean_obj_tag(x_1179) == 4) { -lean_object* x_1187; lean_object* x_1188; lean_object* x_1189; lean_object* x_1190; lean_object* x_1191; lean_object* x_1192; lean_object* x_1193; -x_1187 = lean_unsigned_to_nat(2u); -x_1188 = lean_array_fget(x_907, x_1187); -x_1189 = lean_unsigned_to_nat(3u); -x_1190 = lean_array_fget(x_907, x_1189); -lean_dec(x_907); -x_1191 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__62; -x_1192 = l_Lean_mkApp4(x_1191, x_1180, x_1188, x_1190, x_3); -x_1193 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_1192, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_1193; -} -else +lean_object* x_1180; +x_1180 = lean_ctor_get(x_1179, 0); +lean_inc(x_1180); +if (lean_obj_tag(x_1180) == 1) { -lean_object* x_1194; lean_object* x_1195; -lean_dec(x_1181); +lean_object* x_1181; +x_1181 = lean_ctor_get(x_1180, 0); +lean_inc(x_1181); +if (lean_obj_tag(x_1181) == 0) +{ +lean_object* x_1182; lean_object* x_1183; lean_object* x_1184; lean_object* x_1185; uint8_t x_1186; +x_1182 = lean_ctor_get(x_1146, 1); +lean_inc(x_1182); +lean_dec(x_1146); +x_1183 = lean_ctor_get(x_1179, 1); +lean_inc(x_1183); +lean_dec(x_1179); +x_1184 = lean_ctor_get(x_1180, 1); +lean_inc(x_1184); lean_dec(x_1180); -lean_dec(x_907); +x_1185 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__4; +x_1186 = lean_string_dec_eq(x_1184, x_1185); +lean_dec(x_1184); +if (x_1186 == 0) +{ +lean_object* x_1187; lean_object* x_1188; +lean_dec(x_1183); +lean_dec(x_1182); +lean_dec(x_909); +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_3); -x_1194 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1194, 0, x_2); -lean_ctor_set(x_1194, 1, x_1143); -x_1195 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1195, 0, x_1194); -lean_ctor_set(x_1195, 1, x_13); -return x_1195; -} +x_1187 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1187, 0, x_2); +lean_ctor_set(x_1187, 1, x_1145); +x_1188 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1188, 0, x_1187); +lean_ctor_set(x_1188, 1, x_15); +return x_1188; } +else +{ +if (lean_obj_tag(x_1183) == 0) +{ +lean_object* x_1189; lean_object* x_1190; lean_object* x_1191; lean_object* x_1192; lean_object* x_1193; lean_object* x_1194; lean_object* x_1195; +x_1189 = lean_unsigned_to_nat(2u); +x_1190 = lean_array_fget(x_909, x_1189); +x_1191 = lean_unsigned_to_nat(3u); +x_1192 = lean_array_fget(x_909, x_1191); +lean_dec(x_909); +x_1193 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__62; +x_1194 = l_Lean_mkApp4(x_1193, x_1182, x_1190, x_1192, x_3); +x_1195 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_1194, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_1195; } else { lean_object* x_1196; lean_object* x_1197; -lean_dec(x_1179); -lean_dec(x_1178); -lean_dec(x_1177); -lean_dec(x_1144); -lean_dec(x_907); +lean_dec(x_1183); +lean_dec(x_1182); +lean_dec(x_909); +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_3); x_1196 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1196, 0, x_2); -lean_ctor_set(x_1196, 1, x_1143); +lean_ctor_set(x_1196, 1, x_1145); x_1197 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1197, 0, x_1196); -lean_ctor_set(x_1197, 1, x_13); +lean_ctor_set(x_1197, 1, x_15); return x_1197; } } +} else { lean_object* x_1198; lean_object* x_1199; -lean_dec(x_1178); -lean_dec(x_1177); -lean_dec(x_1144); -lean_dec(x_907); +lean_dec(x_1181); +lean_dec(x_1180); +lean_dec(x_1179); +lean_dec(x_1146); +lean_dec(x_909); +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_3); x_1198 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1198, 0, x_2); -lean_ctor_set(x_1198, 1, x_1143); +lean_ctor_set(x_1198, 1, x_1145); x_1199 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1199, 0, x_1198); -lean_ctor_set(x_1199, 1, x_13); +lean_ctor_set(x_1199, 1, x_15); return x_1199; } } else { lean_object* x_1200; lean_object* x_1201; -lean_dec(x_1177); -lean_dec(x_1144); -lean_dec(x_907); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); +lean_dec(x_1180); +lean_dec(x_1179); +lean_dec(x_1146); +lean_dec(x_909); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_3); x_1200 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1200, 0, x_2); -lean_ctor_set(x_1200, 1, x_1143); +lean_ctor_set(x_1200, 1, x_1145); x_1201 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1201, 0, x_1200); -lean_ctor_set(x_1201, 1, x_13); +lean_ctor_set(x_1201, 1, x_15); return x_1201; } } -default: +else { lean_object* x_1202; lean_object* x_1203; -lean_dec(x_1144); -lean_dec(x_907); +lean_dec(x_1179); +lean_dec(x_1146); +lean_dec(x_909); +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_3); x_1202 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1202, 0, x_2); -lean_ctor_set(x_1202, 1, x_1143); +lean_ctor_set(x_1202, 1, x_1145); x_1203 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1203, 0, x_1202); -lean_ctor_set(x_1203, 1, x_13); +lean_ctor_set(x_1203, 1, x_15); return x_1203; } } +default: +{ +lean_object* x_1204; lean_object* x_1205; +lean_dec(x_1146); +lean_dec(x_909); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_3); +x_1204 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1204, 0, x_2); +lean_ctor_set(x_1204, 1, x_1145); +x_1205 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1205, 0, x_1204); +lean_ctor_set(x_1205, 1, x_15); +return x_1205; +} +} } } } } else { -lean_object* x_1204; uint8_t x_1205; -lean_dec(x_909); -x_1204 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__11; -x_1205 = lean_string_dec_eq(x_908, x_1204); -lean_dec(x_908); -if (x_1205 == 0) +lean_object* x_1206; uint8_t x_1207; +lean_dec(x_911); +x_1206 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__11; +x_1207 = lean_string_dec_eq(x_910, x_1206); +lean_dec(x_910); +if (x_1207 == 0) { -lean_object* x_1206; lean_object* x_1207; lean_object* x_1208; -lean_dec(x_907); +lean_object* x_1208; lean_object* x_1209; lean_object* x_1210; +lean_dec(x_909); +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_3); -x_1206 = lean_unsigned_to_nat(0u); -x_1207 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1207, 0, x_2); -lean_ctor_set(x_1207, 1, x_1206); -x_1208 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1208, 0, x_1207); -lean_ctor_set(x_1208, 1, x_13); -return x_1208; +x_1208 = lean_unsigned_to_nat(0u); +x_1209 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1209, 0, x_2); +lean_ctor_set(x_1209, 1, x_1208); +x_1210 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1210, 0, x_1209); +lean_ctor_set(x_1210, 1, x_15); +return x_1210; } else { -lean_object* x_1209; lean_object* x_1210; uint8_t x_1211; -x_1209 = lean_array_get_size(x_907); -x_1210 = lean_unsigned_to_nat(4u); -x_1211 = lean_nat_dec_eq(x_1209, x_1210); -lean_dec(x_1209); -if (x_1211 == 0) +lean_object* x_1211; lean_object* x_1212; uint8_t x_1213; +x_1211 = lean_array_get_size(x_909); +x_1212 = lean_unsigned_to_nat(4u); +x_1213 = lean_nat_dec_eq(x_1211, x_1212); +lean_dec(x_1211); +if (x_1213 == 0) { -lean_object* x_1212; lean_object* x_1213; lean_object* x_1214; -lean_dec(x_907); +lean_object* x_1214; lean_object* x_1215; lean_object* x_1216; +lean_dec(x_909); +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_3); -x_1212 = lean_unsigned_to_nat(0u); -x_1213 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1213, 0, x_2); -lean_ctor_set(x_1213, 1, x_1212); -x_1214 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1214, 0, x_1213); -lean_ctor_set(x_1214, 1, x_13); -return x_1214; +x_1214 = lean_unsigned_to_nat(0u); +x_1215 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1215, 0, x_2); +lean_ctor_set(x_1215, 1, x_1214); +x_1216 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1216, 0, x_1215); +lean_ctor_set(x_1216, 1, x_15); +return x_1216; } else { -lean_object* x_1215; lean_object* x_1216; -x_1215 = lean_unsigned_to_nat(0u); -x_1216 = lean_array_fget(x_907, x_1215); -switch (lean_obj_tag(x_1216)) { +lean_object* x_1217; lean_object* x_1218; +x_1217 = lean_unsigned_to_nat(0u); +x_1218 = lean_array_fget(x_909, x_1217); +switch (lean_obj_tag(x_1218)) { case 4: { -lean_object* x_1217; -x_1217 = lean_ctor_get(x_1216, 0); -lean_inc(x_1217); -if (lean_obj_tag(x_1217) == 1) -{ -lean_object* x_1218; -x_1218 = lean_ctor_get(x_1217, 0); -lean_inc(x_1218); -if (lean_obj_tag(x_1218) == 0) -{ -lean_object* x_1219; lean_object* x_1220; lean_object* x_1221; uint8_t x_1222; -x_1219 = lean_ctor_get(x_1216, 1); +lean_object* x_1219; +x_1219 = lean_ctor_get(x_1218, 0); lean_inc(x_1219); -lean_dec(x_1216); -x_1220 = lean_ctor_get(x_1217, 1); +if (lean_obj_tag(x_1219) == 1) +{ +lean_object* x_1220; +x_1220 = lean_ctor_get(x_1219, 0); lean_inc(x_1220); -lean_dec(x_1217); -x_1221 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__9; -x_1222 = lean_string_dec_eq(x_1220, x_1221); -if (x_1222 == 0) -{ -lean_object* x_1223; uint8_t x_1224; -x_1223 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__1; -x_1224 = lean_string_dec_eq(x_1220, x_1223); -lean_dec(x_1220); -if (x_1224 == 0) +if (lean_obj_tag(x_1220) == 0) { -lean_object* x_1225; lean_object* x_1226; +lean_object* x_1221; lean_object* x_1222; lean_object* x_1223; uint8_t x_1224; +x_1221 = lean_ctor_get(x_1218, 1); +lean_inc(x_1221); +lean_dec(x_1218); +x_1222 = lean_ctor_get(x_1219, 1); +lean_inc(x_1222); lean_dec(x_1219); -lean_dec(x_907); +x_1223 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__9; +x_1224 = lean_string_dec_eq(x_1222, x_1223); +if (x_1224 == 0) +{ +lean_object* x_1225; uint8_t x_1226; +x_1225 = l_Lean_Elab_Tactic_Omega_succ_x3f___closed__1; +x_1226 = lean_string_dec_eq(x_1222, x_1225); +lean_dec(x_1222); +if (x_1226 == 0) +{ +lean_object* x_1227; lean_object* x_1228; +lean_dec(x_1221); +lean_dec(x_909); +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_3); -x_1225 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1225, 0, x_2); -lean_ctor_set(x_1225, 1, x_1215); -x_1226 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1226, 0, x_1225); -lean_ctor_set(x_1226, 1, x_13); -return x_1226; +x_1227 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1227, 0, x_2); +lean_ctor_set(x_1227, 1, x_1217); +x_1228 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1228, 0, x_1227); +lean_ctor_set(x_1228, 1, x_15); +return x_1228; } else { -if (lean_obj_tag(x_1219) == 0) +if (lean_obj_tag(x_1221) == 0) { -lean_object* x_1227; lean_object* x_1228; lean_object* x_1229; lean_object* x_1230; lean_object* x_1231; lean_object* x_1232; lean_object* x_1233; -x_1227 = lean_unsigned_to_nat(2u); -x_1228 = lean_array_fget(x_907, x_1227); -x_1229 = lean_unsigned_to_nat(3u); -x_1230 = lean_array_fget(x_907, x_1229); -lean_dec(x_907); -x_1231 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__65; -x_1232 = l_Lean_mkApp3(x_1231, x_1228, x_1230, x_3); -x_1233 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_1232, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_1233; +lean_object* x_1229; lean_object* x_1230; lean_object* x_1231; lean_object* x_1232; lean_object* x_1233; lean_object* x_1234; lean_object* x_1235; +x_1229 = lean_unsigned_to_nat(2u); +x_1230 = lean_array_fget(x_909, x_1229); +x_1231 = lean_unsigned_to_nat(3u); +x_1232 = lean_array_fget(x_909, x_1231); +lean_dec(x_909); +x_1233 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__65; +x_1234 = l_Lean_mkApp3(x_1233, x_1230, x_1232, x_3); +x_1235 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_1234, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_1235; } else { -lean_object* x_1234; lean_object* x_1235; -lean_dec(x_1219); -lean_dec(x_907); +lean_object* x_1236; lean_object* x_1237; +lean_dec(x_1221); +lean_dec(x_909); +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_3); -x_1234 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1234, 0, x_2); -lean_ctor_set(x_1234, 1, x_1215); -x_1235 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1235, 0, x_1234); -lean_ctor_set(x_1235, 1, x_13); -return x_1235; +x_1236 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1236, 0, x_2); +lean_ctor_set(x_1236, 1, x_1217); +x_1237 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1237, 0, x_1236); +lean_ctor_set(x_1237, 1, x_15); +return x_1237; } } } else { -lean_dec(x_1220); -if (lean_obj_tag(x_1219) == 0) -{ -lean_object* x_1236; lean_object* x_1237; lean_object* x_1238; lean_object* x_1239; lean_object* x_1240; -x_1236 = lean_unsigned_to_nat(2u); -x_1237 = lean_array_fget(x_907, x_1236); -x_1238 = lean_unsigned_to_nat(3u); -x_1239 = lean_array_fget(x_907, x_1238); -lean_dec(x_907); -lean_inc(x_1237); -x_1240 = l_Lean_Expr_int_x3f(x_1237); -if (lean_obj_tag(x_1240) == 0) -{ -lean_object* x_1241; lean_object* x_1242; lean_object* x_1243; -x_1241 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__68; -x_1242 = l_Lean_mkApp3(x_1241, x_1239, x_1237, x_3); -x_1243 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_1242, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_1243; -} -else -{ -lean_object* x_1244; lean_object* x_1245; uint8_t x_1246; -x_1244 = lean_ctor_get(x_1240, 0); -lean_inc(x_1244); -lean_dec(x_1240); -x_1245 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__8; -x_1246 = lean_int_dec_eq(x_1244, x_1245); -lean_dec(x_1244); -if (x_1246 == 0) -{ -lean_object* x_1247; lean_object* x_1248; lean_object* x_1249; -x_1247 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__68; -x_1248 = l_Lean_mkApp3(x_1247, x_1239, x_1237, x_3); -x_1249 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_1248, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_1249; -} -else -{ -lean_object* x_1250; -lean_dec(x_1237); -x_1250 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality(x_2, x_3, x_1239, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_1250) == 0) -{ -lean_object* x_1251; lean_object* x_1252; lean_object* x_1253; lean_object* x_1254; lean_object* x_1255; lean_object* x_1256; -x_1251 = lean_ctor_get(x_1250, 0); -lean_inc(x_1251); -x_1252 = lean_ctor_get(x_1250, 1); -lean_inc(x_1252); -if (lean_is_exclusive(x_1250)) { - lean_ctor_release(x_1250, 0); - lean_ctor_release(x_1250, 1); - x_1253 = x_1250; +lean_dec(x_1222); +if (lean_obj_tag(x_1221) == 0) +{ +lean_object* x_1238; lean_object* x_1239; lean_object* x_1240; lean_object* x_1241; lean_object* x_1242; +x_1238 = lean_unsigned_to_nat(2u); +x_1239 = lean_array_fget(x_909, x_1238); +x_1240 = lean_unsigned_to_nat(3u); +x_1241 = lean_array_fget(x_909, x_1240); +lean_dec(x_909); +lean_inc(x_1239); +x_1242 = l_Lean_Expr_int_x3f(x_1239); +if (lean_obj_tag(x_1242) == 0) +{ +lean_object* x_1243; lean_object* x_1244; lean_object* x_1245; +x_1243 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__68; +x_1244 = l_Lean_mkApp3(x_1243, x_1241, x_1239, x_3); +x_1245 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_1244, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_1245; +} +else +{ +lean_object* x_1246; lean_object* x_1247; uint8_t x_1248; +x_1246 = lean_ctor_get(x_1242, 0); +lean_inc(x_1246); +lean_dec(x_1242); +x_1247 = l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__8; +x_1248 = lean_int_dec_eq(x_1246, x_1247); +lean_dec(x_1246); +if (x_1248 == 0) +{ +lean_object* x_1249; lean_object* x_1250; lean_object* x_1251; +x_1249 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__68; +x_1250 = l_Lean_mkApp3(x_1249, x_1241, x_1239, x_3); +x_1251 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_1250, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_1251; +} +else +{ +lean_object* x_1252; +lean_dec(x_1239); +x_1252 = l_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality(x_2, x_3, x_1241, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_1252) == 0) +{ +lean_object* x_1253; lean_object* x_1254; lean_object* x_1255; lean_object* x_1256; lean_object* x_1257; lean_object* x_1258; +x_1253 = lean_ctor_get(x_1252, 0); +lean_inc(x_1253); +x_1254 = lean_ctor_get(x_1252, 1); +lean_inc(x_1254); +if (lean_is_exclusive(x_1252)) { + lean_ctor_release(x_1252, 0); + lean_ctor_release(x_1252, 1); + x_1255 = x_1252; } else { - lean_dec_ref(x_1250); - x_1253 = lean_box(0); -} -x_1254 = lean_unsigned_to_nat(1u); -x_1255 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1255, 0, x_1251); -lean_ctor_set(x_1255, 1, x_1254); -if (lean_is_scalar(x_1253)) { - x_1256 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_1252); + x_1255 = lean_box(0); +} +x_1256 = lean_unsigned_to_nat(1u); +x_1257 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1257, 0, x_1253); +lean_ctor_set(x_1257, 1, x_1256); +if (lean_is_scalar(x_1255)) { + x_1258 = lean_alloc_ctor(0, 2, 0); } else { - x_1256 = x_1253; + x_1258 = x_1255; } -lean_ctor_set(x_1256, 0, x_1255); -lean_ctor_set(x_1256, 1, x_1252); -return x_1256; +lean_ctor_set(x_1258, 0, x_1257); +lean_ctor_set(x_1258, 1, x_1254); +return x_1258; } else { -lean_object* x_1257; lean_object* x_1258; lean_object* x_1259; lean_object* x_1260; -x_1257 = lean_ctor_get(x_1250, 0); -lean_inc(x_1257); -x_1258 = lean_ctor_get(x_1250, 1); -lean_inc(x_1258); -if (lean_is_exclusive(x_1250)) { - lean_ctor_release(x_1250, 0); - lean_ctor_release(x_1250, 1); - x_1259 = x_1250; +lean_object* x_1259; lean_object* x_1260; lean_object* x_1261; lean_object* x_1262; +x_1259 = lean_ctor_get(x_1252, 0); +lean_inc(x_1259); +x_1260 = lean_ctor_get(x_1252, 1); +lean_inc(x_1260); +if (lean_is_exclusive(x_1252)) { + lean_ctor_release(x_1252, 0); + lean_ctor_release(x_1252, 1); + x_1261 = x_1252; } else { - lean_dec_ref(x_1250); - x_1259 = lean_box(0); + lean_dec_ref(x_1252); + x_1261 = lean_box(0); } -if (lean_is_scalar(x_1259)) { - x_1260 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_1261)) { + x_1262 = lean_alloc_ctor(1, 2, 0); } else { - x_1260 = x_1259; + x_1262 = x_1261; } -lean_ctor_set(x_1260, 0, x_1257); -lean_ctor_set(x_1260, 1, x_1258); -return x_1260; +lean_ctor_set(x_1262, 0, x_1259); +lean_ctor_set(x_1262, 1, x_1260); +return x_1262; } } } } else { -lean_object* x_1261; lean_object* x_1262; -lean_dec(x_1219); -lean_dec(x_907); +lean_object* x_1263; lean_object* x_1264; +lean_dec(x_1221); +lean_dec(x_909); +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_3); -x_1261 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1261, 0, x_2); -lean_ctor_set(x_1261, 1, x_1215); -x_1262 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1262, 0, x_1261); -lean_ctor_set(x_1262, 1, x_13); -return x_1262; +x_1263 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1263, 0, x_2); +lean_ctor_set(x_1263, 1, x_1217); +x_1264 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1264, 0, x_1263); +lean_ctor_set(x_1264, 1, x_15); +return x_1264; } } } else { -lean_object* x_1263; lean_object* x_1264; +lean_object* x_1265; lean_object* x_1266; +lean_dec(x_1220); +lean_dec(x_1219); lean_dec(x_1218); -lean_dec(x_1217); -lean_dec(x_1216); -lean_dec(x_907); +lean_dec(x_909); +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_3); -x_1263 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1263, 0, x_2); -lean_ctor_set(x_1263, 1, x_1215); -x_1264 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1264, 0, x_1263); -lean_ctor_set(x_1264, 1, x_13); -return x_1264; +x_1265 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1265, 0, x_2); +lean_ctor_set(x_1265, 1, x_1217); +x_1266 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1266, 0, x_1265); +lean_ctor_set(x_1266, 1, x_15); +return x_1266; } } else { -lean_object* x_1265; lean_object* x_1266; -lean_dec(x_1217); -lean_dec(x_1216); -lean_dec(x_907); +lean_object* x_1267; lean_object* x_1268; +lean_dec(x_1219); +lean_dec(x_1218); +lean_dec(x_909); +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_3); -x_1265 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1265, 0, x_2); -lean_ctor_set(x_1265, 1, x_1215); -x_1266 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1266, 0, x_1265); -lean_ctor_set(x_1266, 1, x_13); -return x_1266; +x_1267 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1267, 0, x_2); +lean_ctor_set(x_1267, 1, x_1217); +x_1268 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1268, 0, x_1267); +lean_ctor_set(x_1268, 1, x_15); +return x_1268; } } case 5: { -lean_object* x_1267; -x_1267 = lean_ctor_get(x_1216, 0); -lean_inc(x_1267); -if (lean_obj_tag(x_1267) == 4) -{ -lean_object* x_1268; -x_1268 = lean_ctor_get(x_1267, 0); -lean_inc(x_1268); -if (lean_obj_tag(x_1268) == 1) -{ lean_object* x_1269; -x_1269 = lean_ctor_get(x_1268, 0); +x_1269 = lean_ctor_get(x_1218, 0); lean_inc(x_1269); -if (lean_obj_tag(x_1269) == 0) +if (lean_obj_tag(x_1269) == 4) { -lean_object* x_1270; lean_object* x_1271; lean_object* x_1272; lean_object* x_1273; uint8_t x_1274; -x_1270 = lean_ctor_get(x_1216, 1); +lean_object* x_1270; +x_1270 = lean_ctor_get(x_1269, 0); lean_inc(x_1270); -lean_dec(x_1216); -x_1271 = lean_ctor_get(x_1267, 1); +if (lean_obj_tag(x_1270) == 1) +{ +lean_object* x_1271; +x_1271 = lean_ctor_get(x_1270, 0); lean_inc(x_1271); -lean_dec(x_1267); -x_1272 = lean_ctor_get(x_1268, 1); -lean_inc(x_1272); -lean_dec(x_1268); -x_1273 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__4; -x_1274 = lean_string_dec_eq(x_1272, x_1273); -lean_dec(x_1272); -if (x_1274 == 0) +if (lean_obj_tag(x_1271) == 0) { -lean_object* x_1275; lean_object* x_1276; -lean_dec(x_1271); +lean_object* x_1272; lean_object* x_1273; lean_object* x_1274; lean_object* x_1275; uint8_t x_1276; +x_1272 = lean_ctor_get(x_1218, 1); +lean_inc(x_1272); +lean_dec(x_1218); +x_1273 = lean_ctor_get(x_1269, 1); +lean_inc(x_1273); +lean_dec(x_1269); +x_1274 = lean_ctor_get(x_1270, 1); +lean_inc(x_1274); lean_dec(x_1270); -lean_dec(x_907); +x_1275 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__4; +x_1276 = lean_string_dec_eq(x_1274, x_1275); +lean_dec(x_1274); +if (x_1276 == 0) +{ +lean_object* x_1277; lean_object* x_1278; +lean_dec(x_1273); +lean_dec(x_1272); +lean_dec(x_909); +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_3); -x_1275 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1275, 0, x_2); -lean_ctor_set(x_1275, 1, x_1215); -x_1276 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1276, 0, x_1275); -lean_ctor_set(x_1276, 1, x_13); -return x_1276; +x_1277 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1277, 0, x_2); +lean_ctor_set(x_1277, 1, x_1217); +x_1278 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1278, 0, x_1277); +lean_ctor_set(x_1278, 1, x_15); +return x_1278; } else { -if (lean_obj_tag(x_1271) == 0) -{ -lean_object* x_1277; lean_object* x_1278; lean_object* x_1279; lean_object* x_1280; lean_object* x_1281; lean_object* x_1282; lean_object* x_1283; -x_1277 = lean_unsigned_to_nat(2u); -x_1278 = lean_array_fget(x_907, x_1277); -x_1279 = lean_unsigned_to_nat(3u); -x_1280 = lean_array_fget(x_907, x_1279); -lean_dec(x_907); -x_1281 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__71; -x_1282 = l_Lean_mkApp4(x_1281, x_1270, x_1278, x_1280, x_3); -x_1283 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_1282, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_1283; -} -else +if (lean_obj_tag(x_1273) == 0) { -lean_object* x_1284; lean_object* x_1285; -lean_dec(x_1271); -lean_dec(x_1270); -lean_dec(x_907); -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_3); -x_1284 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1284, 0, x_2); -lean_ctor_set(x_1284, 1, x_1215); -x_1285 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1285, 0, x_1284); -lean_ctor_set(x_1285, 1, x_13); +lean_object* x_1279; lean_object* x_1280; lean_object* x_1281; lean_object* x_1282; lean_object* x_1283; lean_object* x_1284; lean_object* x_1285; +x_1279 = lean_unsigned_to_nat(2u); +x_1280 = lean_array_fget(x_909, x_1279); +x_1281 = lean_unsigned_to_nat(3u); +x_1282 = lean_array_fget(x_909, x_1281); +lean_dec(x_909); +x_1283 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__71; +x_1284 = l_Lean_mkApp4(x_1283, x_1272, x_1280, x_1282, x_3); +x_1285 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_1284, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); return x_1285; } -} -} else { lean_object* x_1286; lean_object* x_1287; -lean_dec(x_1269); -lean_dec(x_1268); -lean_dec(x_1267); -lean_dec(x_1216); -lean_dec(x_907); +lean_dec(x_1273); +lean_dec(x_1272); +lean_dec(x_909); +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_3); x_1286 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1286, 0, x_2); -lean_ctor_set(x_1286, 1, x_1215); +lean_ctor_set(x_1286, 1, x_1217); x_1287 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1287, 0, x_1286); -lean_ctor_set(x_1287, 1, x_13); +lean_ctor_set(x_1287, 1, x_15); return x_1287; } } +} else { lean_object* x_1288; lean_object* x_1289; -lean_dec(x_1268); -lean_dec(x_1267); -lean_dec(x_1216); -lean_dec(x_907); +lean_dec(x_1271); +lean_dec(x_1270); +lean_dec(x_1269); +lean_dec(x_1218); +lean_dec(x_909); +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_3); x_1288 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1288, 0, x_2); -lean_ctor_set(x_1288, 1, x_1215); +lean_ctor_set(x_1288, 1, x_1217); x_1289 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1289, 0, x_1288); -lean_ctor_set(x_1289, 1, x_13); +lean_ctor_set(x_1289, 1, x_15); return x_1289; } } else { lean_object* x_1290; lean_object* x_1291; -lean_dec(x_1267); -lean_dec(x_1216); -lean_dec(x_907); +lean_dec(x_1270); +lean_dec(x_1269); +lean_dec(x_1218); +lean_dec(x_909); +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_3); x_1290 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1290, 0, x_2); -lean_ctor_set(x_1290, 1, x_1215); +lean_ctor_set(x_1290, 1, x_1217); x_1291 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1291, 0, x_1290); -lean_ctor_set(x_1291, 1, x_13); +lean_ctor_set(x_1291, 1, x_15); return x_1291; } } -default: +else { lean_object* x_1292; lean_object* x_1293; -lean_dec(x_1216); -lean_dec(x_907); +lean_dec(x_1269); +lean_dec(x_1218); +lean_dec(x_909); +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_3); x_1292 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1292, 0, x_2); -lean_ctor_set(x_1292, 1, x_1215); +lean_ctor_set(x_1292, 1, x_1217); x_1293 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1293, 0, x_1292); -lean_ctor_set(x_1293, 1, x_13); +lean_ctor_set(x_1293, 1, x_15); return x_1293; } } +default: +{ +lean_object* x_1294; lean_object* x_1295; +lean_dec(x_1218); +lean_dec(x_909); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_3); +x_1294 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1294, 0, x_2); +lean_ctor_set(x_1294, 1, x_1217); +x_1295 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1295, 0, x_1294); +lean_ctor_set(x_1295, 1, x_15); +return x_1295; +} +} } } } @@ -26490,370 +26963,378 @@ return x_1293; } else { -uint8_t x_1294; -lean_dec(x_578); -lean_dec(x_16); -lean_dec(x_15); +uint8_t x_1296; +lean_dec(x_580); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_14); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_3); -x_1294 = !lean_is_exclusive(x_14); -if (x_1294 == 0) +x_1296 = !lean_is_exclusive(x_16); +if (x_1296 == 0) { -lean_object* x_1295; lean_object* x_1296; lean_object* x_1297; lean_object* x_1298; -x_1295 = lean_ctor_get(x_14, 1); -lean_dec(x_1295); -x_1296 = lean_ctor_get(x_14, 0); -lean_dec(x_1296); -x_1297 = lean_unsigned_to_nat(0u); -lean_ctor_set(x_14, 1, x_1297); -lean_ctor_set(x_14, 0, x_2); -x_1298 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1298, 0, x_14); -lean_ctor_set(x_1298, 1, x_13); -return x_1298; +lean_object* x_1297; lean_object* x_1298; lean_object* x_1299; lean_object* x_1300; +x_1297 = lean_ctor_get(x_16, 1); +lean_dec(x_1297); +x_1298 = lean_ctor_get(x_16, 0); +lean_dec(x_1298); +x_1299 = lean_unsigned_to_nat(0u); +lean_ctor_set(x_16, 1, x_1299); +lean_ctor_set(x_16, 0, x_2); +x_1300 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1300, 0, x_16); +lean_ctor_set(x_1300, 1, x_15); +return x_1300; } else { -lean_object* x_1299; lean_object* x_1300; lean_object* x_1301; -lean_dec(x_14); -x_1299 = lean_unsigned_to_nat(0u); -x_1300 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1300, 0, x_2); -lean_ctor_set(x_1300, 1, x_1299); -x_1301 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1301, 0, x_1300); -lean_ctor_set(x_1301, 1, x_13); -return x_1301; +lean_object* x_1301; lean_object* x_1302; lean_object* x_1303; +lean_dec(x_16); +x_1301 = lean_unsigned_to_nat(0u); +x_1302 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1302, 0, x_2); +lean_ctor_set(x_1302, 1, x_1301); +x_1303 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1303, 0, x_1302); +lean_ctor_set(x_1303, 1, x_15); +return x_1303; } } } default: { -uint8_t x_1302; -lean_dec(x_16); -lean_dec(x_15); +uint8_t x_1304; +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_14); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); -x_1302 = !lean_is_exclusive(x_14); -if (x_1302 == 0) +x_1304 = !lean_is_exclusive(x_16); +if (x_1304 == 0) { -lean_object* x_1303; lean_object* x_1304; lean_object* x_1305; lean_object* x_1306; -x_1303 = lean_ctor_get(x_14, 1); -lean_dec(x_1303); -x_1304 = lean_ctor_get(x_14, 0); -lean_dec(x_1304); -x_1305 = lean_unsigned_to_nat(0u); -lean_ctor_set(x_14, 1, x_1305); -lean_ctor_set(x_14, 0, x_2); -x_1306 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1306, 0, x_14); -lean_ctor_set(x_1306, 1, x_13); -return x_1306; +lean_object* x_1305; lean_object* x_1306; lean_object* x_1307; lean_object* x_1308; +x_1305 = lean_ctor_get(x_16, 1); +lean_dec(x_1305); +x_1306 = lean_ctor_get(x_16, 0); +lean_dec(x_1306); +x_1307 = lean_unsigned_to_nat(0u); +lean_ctor_set(x_16, 1, x_1307); +lean_ctor_set(x_16, 0, x_2); +x_1308 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1308, 0, x_16); +lean_ctor_set(x_1308, 1, x_15); +return x_1308; } else { -lean_object* x_1307; lean_object* x_1308; lean_object* x_1309; -lean_dec(x_14); -x_1307 = lean_unsigned_to_nat(0u); -x_1308 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1308, 0, x_2); -lean_ctor_set(x_1308, 1, x_1307); -x_1309 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1309, 0, x_1308); -lean_ctor_set(x_1309, 1, x_13); -return x_1309; +lean_object* x_1309; lean_object* x_1310; lean_object* x_1311; +lean_dec(x_16); +x_1309 = lean_unsigned_to_nat(0u); +x_1310 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1310, 0, x_2); +lean_ctor_set(x_1310, 1, x_1309); +x_1311 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1311, 0, x_1310); +lean_ctor_set(x_1311, 1, x_15); +return x_1311; } } } } else { -uint8_t x_1310; -lean_dec(x_15); +uint8_t x_1312; +lean_dec(x_17); +lean_dec(x_14); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); -x_1310 = !lean_is_exclusive(x_14); -if (x_1310 == 0) +x_1312 = !lean_is_exclusive(x_16); +if (x_1312 == 0) { -lean_object* x_1311; lean_object* x_1312; lean_object* x_1313; lean_object* x_1314; -x_1311 = lean_ctor_get(x_14, 1); -lean_dec(x_1311); -x_1312 = lean_ctor_get(x_14, 0); -lean_dec(x_1312); -x_1313 = lean_unsigned_to_nat(0u); -lean_ctor_set(x_14, 1, x_1313); -lean_ctor_set(x_14, 0, x_2); -x_1314 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1314, 0, x_14); -lean_ctor_set(x_1314, 1, x_13); -return x_1314; +lean_object* x_1313; lean_object* x_1314; lean_object* x_1315; lean_object* x_1316; +x_1313 = lean_ctor_get(x_16, 1); +lean_dec(x_1313); +x_1314 = lean_ctor_get(x_16, 0); +lean_dec(x_1314); +x_1315 = lean_unsigned_to_nat(0u); +lean_ctor_set(x_16, 1, x_1315); +lean_ctor_set(x_16, 0, x_2); +x_1316 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1316, 0, x_16); +lean_ctor_set(x_1316, 1, x_15); +return x_1316; } else { -lean_object* x_1315; lean_object* x_1316; lean_object* x_1317; -lean_dec(x_14); -x_1315 = lean_unsigned_to_nat(0u); -x_1316 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1316, 0, x_2); -lean_ctor_set(x_1316, 1, x_1315); -x_1317 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1317, 0, x_1316); -lean_ctor_set(x_1317, 1, x_13); -return x_1317; +lean_object* x_1317; lean_object* x_1318; lean_object* x_1319; +lean_dec(x_16); +x_1317 = lean_unsigned_to_nat(0u); +x_1318 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1318, 0, x_2); +lean_ctor_set(x_1318, 1, x_1317); +x_1319 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1319, 0, x_1318); +lean_ctor_set(x_1319, 1, x_15); +return x_1319; } } } case 7: { -lean_object* x_1318; lean_object* x_1319; lean_object* x_1320; +lean_object* x_1320; lean_object* x_1321; lean_object* x_1322; lean_dec(x_4); -x_1318 = lean_ctor_get(x_1, 1); -lean_inc(x_1318); -x_1319 = lean_ctor_get(x_1, 2); -lean_inc(x_1319); +x_1320 = lean_ctor_get(x_1, 1); +lean_inc(x_1320); +x_1321 = lean_ctor_get(x_1, 2); +lean_inc(x_1321); lean_dec(x_1); +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_1318); -x_1320 = l_Lean_Meta_isProp(x_1318, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_1320) == 0) -{ -lean_object* x_1321; lean_object* x_1322; lean_object* x_1323; -x_1321 = lean_ctor_get(x_1320, 0); -lean_inc(x_1321); -x_1322 = lean_ctor_get(x_1320, 1); -lean_inc(x_1322); -lean_dec(x_1320); +lean_inc(x_1320); +x_1322 = l_Lean_Meta_isProp(x_1320, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_1322) == 0) +{ +lean_object* x_1323; lean_object* x_1324; lean_object* x_1325; +x_1323 = lean_ctor_get(x_1322, 0); +lean_inc(x_1323); +x_1324 = lean_ctor_get(x_1322, 1); +lean_inc(x_1324); +lean_dec(x_1322); +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_1319); -x_1323 = l_Lean_Meta_isProp(x_1319, x_9, x_10, x_11, x_12, x_1322); -if (lean_obj_tag(x_1323) == 0) +lean_inc(x_1321); +x_1325 = l_Lean_Meta_isProp(x_1321, x_11, x_12, x_13, x_14, x_1324); +if (lean_obj_tag(x_1325) == 0) { -uint8_t x_1324; -x_1324 = lean_unbox(x_1321); -lean_dec(x_1321); -if (x_1324 == 0) +uint8_t x_1326; +x_1326 = lean_unbox(x_1323); +lean_dec(x_1323); +if (x_1326 == 0) { -uint8_t x_1325; -lean_dec(x_1319); -lean_dec(x_1318); +uint8_t x_1327; +lean_dec(x_1321); +lean_dec(x_1320); +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_3); -x_1325 = !lean_is_exclusive(x_1323); -if (x_1325 == 0) +x_1327 = !lean_is_exclusive(x_1325); +if (x_1327 == 0) { -lean_object* x_1326; lean_object* x_1327; lean_object* x_1328; -x_1326 = lean_ctor_get(x_1323, 0); -lean_dec(x_1326); -x_1327 = lean_unsigned_to_nat(0u); -x_1328 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1328, 0, x_2); -lean_ctor_set(x_1328, 1, x_1327); -lean_ctor_set(x_1323, 0, x_1328); -return x_1323; +lean_object* x_1328; lean_object* x_1329; lean_object* x_1330; +x_1328 = lean_ctor_get(x_1325, 0); +lean_dec(x_1328); +x_1329 = lean_unsigned_to_nat(0u); +x_1330 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1330, 0, x_2); +lean_ctor_set(x_1330, 1, x_1329); +lean_ctor_set(x_1325, 0, x_1330); +return x_1325; } else { -lean_object* x_1329; lean_object* x_1330; lean_object* x_1331; lean_object* x_1332; -x_1329 = lean_ctor_get(x_1323, 1); -lean_inc(x_1329); -lean_dec(x_1323); -x_1330 = lean_unsigned_to_nat(0u); -x_1331 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1331, 0, x_2); -lean_ctor_set(x_1331, 1, x_1330); -x_1332 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1332, 0, x_1331); -lean_ctor_set(x_1332, 1, x_1329); -return x_1332; +lean_object* x_1331; lean_object* x_1332; lean_object* x_1333; lean_object* x_1334; +x_1331 = lean_ctor_get(x_1325, 1); +lean_inc(x_1331); +lean_dec(x_1325); +x_1332 = lean_unsigned_to_nat(0u); +x_1333 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1333, 0, x_2); +lean_ctor_set(x_1333, 1, x_1332); +x_1334 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1334, 0, x_1333); +lean_ctor_set(x_1334, 1, x_1331); +return x_1334; } } else { -lean_object* x_1333; uint8_t x_1334; -x_1333 = lean_ctor_get(x_1323, 0); -lean_inc(x_1333); -x_1334 = lean_unbox(x_1333); -lean_dec(x_1333); -if (x_1334 == 0) +lean_object* x_1335; uint8_t x_1336; +x_1335 = lean_ctor_get(x_1325, 0); +lean_inc(x_1335); +x_1336 = lean_unbox(x_1335); +lean_dec(x_1335); +if (x_1336 == 0) { -uint8_t x_1335; -lean_dec(x_1319); -lean_dec(x_1318); +uint8_t x_1337; +lean_dec(x_1321); +lean_dec(x_1320); +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_3); -x_1335 = !lean_is_exclusive(x_1323); -if (x_1335 == 0) +x_1337 = !lean_is_exclusive(x_1325); +if (x_1337 == 0) { -lean_object* x_1336; lean_object* x_1337; lean_object* x_1338; -x_1336 = lean_ctor_get(x_1323, 0); -lean_dec(x_1336); -x_1337 = lean_unsigned_to_nat(0u); -x_1338 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1338, 0, x_2); -lean_ctor_set(x_1338, 1, x_1337); -lean_ctor_set(x_1323, 0, x_1338); -return x_1323; +lean_object* x_1338; lean_object* x_1339; lean_object* x_1340; +x_1338 = lean_ctor_get(x_1325, 0); +lean_dec(x_1338); +x_1339 = lean_unsigned_to_nat(0u); +x_1340 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1340, 0, x_2); +lean_ctor_set(x_1340, 1, x_1339); +lean_ctor_set(x_1325, 0, x_1340); +return x_1325; } else { -lean_object* x_1339; lean_object* x_1340; lean_object* x_1341; lean_object* x_1342; -x_1339 = lean_ctor_get(x_1323, 1); -lean_inc(x_1339); -lean_dec(x_1323); -x_1340 = lean_unsigned_to_nat(0u); -x_1341 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1341, 0, x_2); -lean_ctor_set(x_1341, 1, x_1340); -x_1342 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1342, 0, x_1341); -lean_ctor_set(x_1342, 1, x_1339); -return x_1342; +lean_object* x_1341; lean_object* x_1342; lean_object* x_1343; lean_object* x_1344; +x_1341 = lean_ctor_get(x_1325, 1); +lean_inc(x_1341); +lean_dec(x_1325); +x_1342 = lean_unsigned_to_nat(0u); +x_1343 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1343, 0, x_2); +lean_ctor_set(x_1343, 1, x_1342); +x_1344 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1344, 0, x_1343); +lean_ctor_set(x_1344, 1, x_1341); +return x_1344; } } else { -lean_object* x_1343; lean_object* x_1344; lean_object* x_1345; lean_object* x_1346; lean_object* x_1347; lean_object* x_1348; -x_1343 = lean_ctor_get(x_1323, 1); -lean_inc(x_1343); -lean_dec(x_1323); -x_1344 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__1___closed__8; -lean_inc(x_1318); -x_1345 = l_Lean_Expr_app___override(x_1344, x_1318); -x_1346 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__74; -x_1347 = l_Lean_mkApp4(x_1346, x_1318, x_1319, x_1345, x_3); -x_1348 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_1347, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_1343); -return x_1348; +lean_object* x_1345; lean_object* x_1346; lean_object* x_1347; lean_object* x_1348; lean_object* x_1349; lean_object* x_1350; +x_1345 = lean_ctor_get(x_1325, 1); +lean_inc(x_1345); +lean_dec(x_1325); +x_1346 = l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__1___closed__8; +lean_inc(x_1320); +x_1347 = l_Lean_Expr_app___override(x_1346, x_1320); +x_1348 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__74; +x_1349 = l_Lean_mkApp4(x_1348, x_1320, x_1321, x_1347, x_3); +x_1350 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_2, x_1349, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_1345); +return x_1350; } } } else { -uint8_t x_1349; +uint8_t x_1351; +lean_dec(x_1323); lean_dec(x_1321); -lean_dec(x_1319); -lean_dec(x_1318); +lean_dec(x_1320); +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_3); lean_dec(x_2); -x_1349 = !lean_is_exclusive(x_1323); -if (x_1349 == 0) +x_1351 = !lean_is_exclusive(x_1325); +if (x_1351 == 0) { -return x_1323; +return x_1325; } else { -lean_object* x_1350; lean_object* x_1351; lean_object* x_1352; -x_1350 = lean_ctor_get(x_1323, 0); -x_1351 = lean_ctor_get(x_1323, 1); -lean_inc(x_1351); -lean_inc(x_1350); -lean_dec(x_1323); -x_1352 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_1352, 0, x_1350); -lean_ctor_set(x_1352, 1, x_1351); -return x_1352; +lean_object* x_1352; lean_object* x_1353; lean_object* x_1354; +x_1352 = lean_ctor_get(x_1325, 0); +x_1353 = lean_ctor_get(x_1325, 1); +lean_inc(x_1353); +lean_inc(x_1352); +lean_dec(x_1325); +x_1354 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1354, 0, x_1352); +lean_ctor_set(x_1354, 1, x_1353); +return x_1354; } } } else { -uint8_t x_1353; -lean_dec(x_1319); -lean_dec(x_1318); +uint8_t x_1355; +lean_dec(x_1321); +lean_dec(x_1320); +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_3); lean_dec(x_2); -x_1353 = !lean_is_exclusive(x_1320); -if (x_1353 == 0) +x_1355 = !lean_is_exclusive(x_1322); +if (x_1355 == 0) { -return x_1320; +return x_1322; } else { -lean_object* x_1354; lean_object* x_1355; lean_object* x_1356; -x_1354 = lean_ctor_get(x_1320, 0); -x_1355 = lean_ctor_get(x_1320, 1); -lean_inc(x_1355); -lean_inc(x_1354); -lean_dec(x_1320); -x_1356 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_1356, 0, x_1354); -lean_ctor_set(x_1356, 1, x_1355); -return x_1356; +lean_object* x_1356; lean_object* x_1357; lean_object* x_1358; +x_1356 = lean_ctor_get(x_1322, 0); +x_1357 = lean_ctor_get(x_1322, 1); +lean_inc(x_1357); +lean_inc(x_1356); +lean_dec(x_1322); +x_1358 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1358, 0, x_1356); +lean_ctor_set(x_1358, 1, x_1357); +return x_1358; } } } default: { -lean_object* x_1357; lean_object* x_1358; lean_object* x_1359; +lean_object* x_1359; lean_object* x_1360; lean_object* x_1361; +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_4); lean_dec(x_3); lean_dec(x_1); -x_1357 = lean_unsigned_to_nat(0u); -x_1358 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1358, 0, x_2); -lean_ctor_set(x_1358, 1, x_1357); -x_1359 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1359, 0, x_1358); -lean_ctor_set(x_1359, 1, x_13); -return x_1359; +x_1359 = lean_unsigned_to_nat(0u); +x_1360 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1360, 0, x_2); +lean_ctor_set(x_1360, 1, x_1359); +x_1361 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1361, 0, x_1360); +lean_ctor_set(x_1361, 1, x_15); +return x_1361; } } } @@ -26875,571 +27356,612 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_Tactic_Omega_MetaProblem_addFact(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_object* x_11, lean_object* x_12) { _start: { -lean_object* x_11; uint8_t x_12; -x_11 = lean_ctor_get(x_1, 0); -lean_inc(x_11); -x_12 = lean_ctor_get_uint8(x_11, sizeof(void*)*7); -if (x_12 == 0) +lean_object* x_13; uint8_t x_14; +x_13 = lean_ctor_get(x_1, 0); +lean_inc(x_13); +x_14 = lean_ctor_get_uint8(x_13, sizeof(void*)*7); +if (x_14 == 0) { -lean_object* x_13; lean_object* x_14; lean_object* x_15; +lean_object* x_15; lean_object* x_16; lean_object* x_17; +lean_dec(x_13); lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_13 = lean_unsigned_to_nat(0u); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_1); -lean_ctor_set(x_14, 1, x_13); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_14); -lean_ctor_set(x_15, 1, x_10); -return x_15; +x_15 = lean_unsigned_to_nat(0u); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_1); +lean_ctor_set(x_16, 1, 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_12); +return x_17; } else { -lean_object* x_16; +lean_object* x_18; +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); lean_inc(x_2); -x_16 = lean_infer_type(x_2, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_16) == 0) +x_18 = lean_infer_type(x_2, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); -lean_inc(x_18); -lean_dec(x_16); +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_19 = l_Lean_Meta_whnfR(x_17, x_6, x_7, x_8, x_9, x_18); -if (lean_obj_tag(x_19) == 0) +x_21 = l_Lean_Meta_whnfR(x_19, x_8, x_9, x_10, x_11, x_20); +if (lean_obj_tag(x_21) == 0) { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; -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_instantiateMVars___at_Lean_Elab_Tactic_Omega_MetaProblem_addFact___spec__1(x_20, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_21); -x_23 = lean_ctor_get(x_22, 0); +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_21, 1); lean_inc(x_23); -x_24 = lean_ctor_get(x_22, 1); -lean_inc(x_24); -lean_dec(x_22); -x_25 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__2; -x_26 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_25, x_3, x_4, x_5, x_6, x_7, x_8, x_9, 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) -{ -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_26, 1); +lean_dec(x_21); +x_24 = l_Lean_instantiateMVars___at_Lean_Elab_Tactic_Omega_MetaProblem_addFact___spec__1(x_22, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, 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 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__2; +x_28 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_27, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_26); +x_29 = lean_ctor_get(x_28, 0); lean_inc(x_29); -lean_dec(x_26); -x_30 = lean_box(0); -x_31 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1(x_23, x_1, x_2, x_11, x_30, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_29); -return x_31; +x_30 = lean_unbox(x_29); +lean_dec(x_29); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_28, 1); +lean_inc(x_31); +lean_dec(x_28); +x_32 = lean_box(0); +x_33 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1(x_25, x_1, x_2, x_13, x_32, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_31); +return x_33; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_32 = lean_ctor_get(x_26, 1); -lean_inc(x_32); -lean_dec(x_26); -lean_inc(x_23); -x_33 = l_Lean_MessageData_ofExpr(x_23); -x_34 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___closed__2; -x_35 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_33); -x_36 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__9; +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_34 = lean_ctor_get(x_28, 1); +lean_inc(x_34); +lean_dec(x_28); +lean_inc(x_25); +x_35 = l_Lean_MessageData_ofExpr(x_25); +x_36 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___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 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_25, x_37, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_32); -x_39 = lean_ctor_get(x_38, 0); -lean_inc(x_39); -x_40 = lean_ctor_get(x_38, 1); -lean_inc(x_40); -lean_dec(x_38); -x_41 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1(x_23, x_1, x_2, x_11, x_39, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_40); -return x_41; +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_35); +x_38 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__9; +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_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_27, x_39, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_34); +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_Tactic_Omega_MetaProblem_addFact___lambda__1(x_25, x_1, x_2, x_13, x_41, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_42); +return x_43; } } else { -uint8_t x_42; +uint8_t x_44; +lean_dec(x_13); lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_42 = !lean_is_exclusive(x_19); -if (x_42 == 0) +x_44 = !lean_is_exclusive(x_21); +if (x_44 == 0) { -return x_19; +return x_21; } else { -lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_43 = lean_ctor_get(x_19, 0); -x_44 = lean_ctor_get(x_19, 1); -lean_inc(x_44); -lean_inc(x_43); -lean_dec(x_19); -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_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_21, 0); +x_46 = lean_ctor_get(x_21, 1); +lean_inc(x_46); +lean_inc(x_45); +lean_dec(x_21); +x_47 = lean_alloc_ctor(1, 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_46; +uint8_t x_48; +lean_dec(x_13); lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_46 = !lean_is_exclusive(x_16); -if (x_46 == 0) +x_48 = !lean_is_exclusive(x_18); +if (x_48 == 0) { -return x_16; +return x_18; } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_47 = lean_ctor_get(x_16, 0); -x_48 = lean_ctor_get(x_16, 1); -lean_inc(x_48); -lean_inc(x_47); -lean_dec(x_16); -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_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_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_instantiateMVars___at_Lean_Elab_Tactic_Omega_MetaProblem_addFact___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_EXPORT lean_object* l_Lean_instantiateMVars___at_Lean_Elab_Tactic_Omega_MetaProblem_addFact___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_10; -x_10 = l_Lean_instantiateMVars___at_Lean_Elab_Tactic_Omega_MetaProblem_addFact___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +uint8_t x_12; lean_object* x_13; +x_12 = lean_unbox(x_5); +lean_dec(x_5); +x_13 = l_Lean_instantiateMVars___at_Lean_Elab_Tactic_Omega_MetaProblem_addFact___spec__1(x_1, x_2, x_3, x_4, x_12, 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_10; +return x_13; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_processFacts(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { -lean_object* x_10; -x_10 = lean_ctor_get(x_1, 1); -lean_inc(x_10); -if (lean_obj_tag(x_10) == 0) +uint8_t x_16; lean_object* x_17; +x_16 = lean_unbox(x_9); +lean_dec(x_9); +x_17 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1(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_13, x_14, x_15); +return x_17; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; +uint8_t x_13; lean_object* x_14; +x_13 = lean_unbox(x_6); +lean_dec(x_6); +x_14 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_1, x_2, x_3, x_4, x_5, x_13, x_7, x_8, x_9, x_10, x_11, x_12); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_processFacts(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) { +_start: +{ +lean_object* x_12; +x_12 = lean_ctor_get(x_1, 1); +lean_inc(x_12); +if (lean_obj_tag(x_12) == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +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_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); -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; +x_13 = lean_unsigned_to_nat(0u); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_1); +lean_ctor_set(x_14, 1, x_13); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_11); +return x_15; } else { -uint8_t x_14; -x_14 = !lean_is_exclusive(x_1); -if (x_14 == 0) +uint8_t x_16; +x_16 = !lean_is_exclusive(x_1); +if (x_16 == 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_1, 3); -x_16 = lean_ctor_get(x_1, 1); -lean_dec(x_16); -x_17 = lean_ctor_get(x_10, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_10, 1); -lean_inc(x_18); -lean_dec(x_10); +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_1, 3); +x_18 = lean_ctor_get(x_1, 1); +lean_dec(x_18); +x_19 = lean_ctor_get(x_12, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_12, 1); +lean_inc(x_20); +lean_dec(x_12); +lean_inc(x_19); lean_inc(x_17); -lean_inc(x_15); -x_19 = l_Lean_HashSetImp_contains___at_Lean_CollectLevelParams_visitExpr___spec__1(x_15, x_17); -if (x_19 == 0) +x_21 = l_Lean_HashSetImp_contains___at_Lean_CollectLevelParams_visitExpr___spec__1(x_17, x_19); +if (x_21 == 0) { -lean_object* x_20; lean_object* x_21; -lean_inc(x_17); -x_20 = l_Lean_HashSetImp_insert___at_Lean_CollectLevelParams_visitExpr___spec__3(x_15, x_17); -lean_ctor_set(x_1, 3, x_20); -lean_ctor_set(x_1, 1, x_18); +lean_object* x_22; lean_object* x_23; +lean_inc(x_19); +x_22 = l_Lean_HashSetImp_insert___at_Lean_CollectLevelParams_visitExpr___spec__3(x_17, x_19); +lean_ctor_set(x_1, 3, x_22); +lean_ctor_set(x_1, 1, x_20); +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_2); -x_21 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_1, x_17, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_21) == 0) +x_23 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_1, x_19, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_23) == 0) { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -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_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_22, 1); +x_25 = lean_ctor_get(x_23, 1); lean_inc(x_25); -lean_dec(x_22); -x_26 = l_Lean_Elab_Tactic_Omega_MetaProblem_processFacts(x_24, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_23); -if (lean_obj_tag(x_26) == 0) -{ -uint8_t x_27; -x_27 = !lean_is_exclusive(x_26); -if (x_27 == 0) +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); +x_28 = l_Lean_Elab_Tactic_Omega_MetaProblem_processFacts(x_26, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_25); +if (lean_obj_tag(x_28) == 0) { -lean_object* x_28; uint8_t x_29; -x_28 = lean_ctor_get(x_26, 0); +uint8_t x_29; 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, 1); -x_31 = lean_nat_add(x_25, x_30); -lean_dec(x_30); -lean_dec(x_25); -lean_ctor_set(x_28, 1, x_31); -return x_26; +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_nat_add(x_27, x_32); +lean_dec(x_32); +lean_dec(x_27); +lean_ctor_set(x_30, 1, x_33); +return x_28; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_32 = lean_ctor_get(x_28, 0); -x_33 = lean_ctor_get(x_28, 1); -lean_inc(x_33); -lean_inc(x_32); -lean_dec(x_28); -x_34 = lean_nat_add(x_25, x_33); -lean_dec(x_33); -lean_dec(x_25); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_32); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_26, 0, x_35); -return x_26; +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_nat_add(x_27, x_35); +lean_dec(x_35); +lean_dec(x_27); +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_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_36 = lean_ctor_get(x_26, 0); -x_37 = lean_ctor_get(x_26, 1); -lean_inc(x_37); -lean_inc(x_36); -lean_dec(x_26); -x_38 = lean_ctor_get(x_36, 0); -lean_inc(x_38); -x_39 = lean_ctor_get(x_36, 1); +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_38 = lean_ctor_get(x_28, 0); +x_39 = lean_ctor_get(x_28, 1); lean_inc(x_39); -if (lean_is_exclusive(x_36)) { - lean_ctor_release(x_36, 0); - lean_ctor_release(x_36, 1); - x_40 = x_36; +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_36); - x_40 = lean_box(0); + lean_dec_ref(x_38); + x_42 = lean_box(0); } -x_41 = lean_nat_add(x_25, x_39); -lean_dec(x_39); -lean_dec(x_25); -if (lean_is_scalar(x_40)) { - x_42 = lean_alloc_ctor(0, 2, 0); +x_43 = lean_nat_add(x_27, x_41); +lean_dec(x_41); +lean_dec(x_27); +if (lean_is_scalar(x_42)) { + x_44 = lean_alloc_ctor(0, 2, 0); } else { - x_42 = x_40; + x_44 = x_42; } -lean_ctor_set(x_42, 0, x_38); -lean_ctor_set(x_42, 1, 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_37); -return x_43; +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_44; -lean_dec(x_25); -x_44 = !lean_is_exclusive(x_26); -if (x_44 == 0) +uint8_t x_46; +lean_dec(x_27); +x_46 = !lean_is_exclusive(x_28); +if (x_46 == 0) { -return x_26; +return x_28; } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_45 = lean_ctor_get(x_26, 0); -x_46 = lean_ctor_get(x_26, 1); -lean_inc(x_46); -lean_inc(x_45); -lean_dec(x_26); -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_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; } } } else { -uint8_t x_48; +uint8_t x_50; +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_48 = !lean_is_exclusive(x_21); -if (x_48 == 0) +x_50 = !lean_is_exclusive(x_23); +if (x_50 == 0) { -return x_21; +return x_23; } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_49 = lean_ctor_get(x_21, 0); -x_50 = lean_ctor_get(x_21, 1); -lean_inc(x_50); -lean_inc(x_49); -lean_dec(x_21); -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_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_23, 0); +x_52 = lean_ctor_get(x_23, 1); +lean_inc(x_52); +lean_inc(x_51); +lean_dec(x_23); +x_53 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_52); +return x_53; } } } else { -lean_dec(x_17); -lean_ctor_set(x_1, 1, x_18); +lean_dec(x_19); +lean_ctor_set(x_1, 1, x_20); goto _start; } } else { -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_1, 0); -x_54 = lean_ctor_get(x_1, 2); -x_55 = lean_ctor_get(x_1, 3); -lean_inc(x_55); -lean_inc(x_54); -lean_inc(x_53); -lean_dec(x_1); -x_56 = lean_ctor_get(x_10, 0); -lean_inc(x_56); -x_57 = lean_ctor_get(x_10, 1); +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_55 = lean_ctor_get(x_1, 0); +x_56 = lean_ctor_get(x_1, 2); +x_57 = lean_ctor_get(x_1, 3); lean_inc(x_57); -lean_dec(x_10); lean_inc(x_56); lean_inc(x_55); -x_58 = l_Lean_HashSetImp_contains___at_Lean_CollectLevelParams_visitExpr___spec__1(x_55, x_56); -if (x_58 == 0) +lean_dec(x_1); +x_58 = lean_ctor_get(x_12, 0); +lean_inc(x_58); +x_59 = lean_ctor_get(x_12, 1); +lean_inc(x_59); +lean_dec(x_12); +lean_inc(x_58); +lean_inc(x_57); +x_60 = l_Lean_HashSetImp_contains___at_Lean_CollectLevelParams_visitExpr___spec__1(x_57, x_58); +if (x_60 == 0) { -lean_object* x_59; lean_object* x_60; lean_object* x_61; -lean_inc(x_56); -x_59 = l_Lean_HashSetImp_insert___at_Lean_CollectLevelParams_visitExpr___spec__3(x_55, x_56); -x_60 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_60, 0, x_53); -lean_ctor_set(x_60, 1, x_57); -lean_ctor_set(x_60, 2, x_54); -lean_ctor_set(x_60, 3, x_59); +lean_object* x_61; lean_object* x_62; lean_object* x_63; +lean_inc(x_58); +x_61 = l_Lean_HashSetImp_insert___at_Lean_CollectLevelParams_visitExpr___spec__3(x_57, x_58); +x_62 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_62, 0, x_55); +lean_ctor_set(x_62, 1, x_59); +lean_ctor_set(x_62, 2, x_56); +lean_ctor_set(x_62, 3, x_61); +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_2); -x_61 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_60, x_56, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_61) == 0) +x_63 = l_Lean_Elab_Tactic_Omega_MetaProblem_addFact(x_62, x_58, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_63) == 0) { -lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_62 = lean_ctor_get(x_61, 0); -lean_inc(x_62); -x_63 = lean_ctor_get(x_61, 1); -lean_inc(x_63); -lean_dec(x_61); -x_64 = lean_ctor_get(x_62, 0); +lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_64 = lean_ctor_get(x_63, 0); lean_inc(x_64); -x_65 = lean_ctor_get(x_62, 1); +x_65 = lean_ctor_get(x_63, 1); lean_inc(x_65); -lean_dec(x_62); -x_66 = l_Lean_Elab_Tactic_Omega_MetaProblem_processFacts(x_64, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_63); -if (lean_obj_tag(x_66) == 0) -{ -lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_67 = lean_ctor_get(x_66, 0); +lean_dec(x_63); +x_66 = lean_ctor_get(x_64, 0); +lean_inc(x_66); +x_67 = lean_ctor_get(x_64, 1); lean_inc(x_67); -x_68 = lean_ctor_get(x_66, 1); -lean_inc(x_68); -if (lean_is_exclusive(x_66)) { - lean_ctor_release(x_66, 0); - lean_ctor_release(x_66, 1); - x_69 = x_66; +lean_dec(x_64); +x_68 = l_Lean_Elab_Tactic_Omega_MetaProblem_processFacts(x_66, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_65); +if (lean_obj_tag(x_68) == 0) +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_69 = lean_ctor_get(x_68, 0); +lean_inc(x_69); +x_70 = lean_ctor_get(x_68, 1); +lean_inc(x_70); +if (lean_is_exclusive(x_68)) { + lean_ctor_release(x_68, 0); + lean_ctor_release(x_68, 1); + x_71 = x_68; } else { - lean_dec_ref(x_66); - x_69 = lean_box(0); + lean_dec_ref(x_68); + x_71 = lean_box(0); } -x_70 = lean_ctor_get(x_67, 0); -lean_inc(x_70); -x_71 = lean_ctor_get(x_67, 1); -lean_inc(x_71); -if (lean_is_exclusive(x_67)) { - lean_ctor_release(x_67, 0); - lean_ctor_release(x_67, 1); - x_72 = x_67; +x_72 = lean_ctor_get(x_69, 0); +lean_inc(x_72); +x_73 = lean_ctor_get(x_69, 1); +lean_inc(x_73); +if (lean_is_exclusive(x_69)) { + lean_ctor_release(x_69, 0); + lean_ctor_release(x_69, 1); + x_74 = x_69; } else { - lean_dec_ref(x_67); - x_72 = lean_box(0); + lean_dec_ref(x_69); + x_74 = lean_box(0); } -x_73 = lean_nat_add(x_65, x_71); -lean_dec(x_71); -lean_dec(x_65); -if (lean_is_scalar(x_72)) { - x_74 = lean_alloc_ctor(0, 2, 0); +x_75 = lean_nat_add(x_67, x_73); +lean_dec(x_73); +lean_dec(x_67); +if (lean_is_scalar(x_74)) { + x_76 = lean_alloc_ctor(0, 2, 0); } else { - x_74 = x_72; + x_76 = x_74; } -lean_ctor_set(x_74, 0, x_70); -lean_ctor_set(x_74, 1, x_73); -if (lean_is_scalar(x_69)) { - x_75 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_76, 0, x_72); +lean_ctor_set(x_76, 1, x_75); +if (lean_is_scalar(x_71)) { + x_77 = lean_alloc_ctor(0, 2, 0); } else { - x_75 = x_69; + x_77 = x_71; } -lean_ctor_set(x_75, 0, x_74); -lean_ctor_set(x_75, 1, x_68); -return x_75; +lean_ctor_set(x_77, 0, x_76); +lean_ctor_set(x_77, 1, x_70); +return x_77; } else { -lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; -lean_dec(x_65); -x_76 = lean_ctor_get(x_66, 0); -lean_inc(x_76); -x_77 = lean_ctor_get(x_66, 1); -lean_inc(x_77); -if (lean_is_exclusive(x_66)) { - lean_ctor_release(x_66, 0); - lean_ctor_release(x_66, 1); - x_78 = x_66; +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +lean_dec(x_67); +x_78 = lean_ctor_get(x_68, 0); +lean_inc(x_78); +x_79 = lean_ctor_get(x_68, 1); +lean_inc(x_79); +if (lean_is_exclusive(x_68)) { + lean_ctor_release(x_68, 0); + lean_ctor_release(x_68, 1); + x_80 = x_68; } else { - lean_dec_ref(x_66); - x_78 = lean_box(0); + lean_dec_ref(x_68); + x_80 = lean_box(0); } -if (lean_is_scalar(x_78)) { - x_79 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_80)) { + x_81 = lean_alloc_ctor(1, 2, 0); } else { - x_79 = x_78; + x_81 = x_80; } -lean_ctor_set(x_79, 0, x_76); -lean_ctor_set(x_79, 1, x_77); -return x_79; +lean_ctor_set(x_81, 0, x_78); +lean_ctor_set(x_81, 1, x_79); +return x_81; } } 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_6); -lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_80 = lean_ctor_get(x_61, 0); -lean_inc(x_80); -x_81 = lean_ctor_get(x_61, 1); -lean_inc(x_81); -if (lean_is_exclusive(x_61)) { - lean_ctor_release(x_61, 0); - lean_ctor_release(x_61, 1); - x_82 = x_61; +x_82 = lean_ctor_get(x_63, 0); +lean_inc(x_82); +x_83 = lean_ctor_get(x_63, 1); +lean_inc(x_83); +if (lean_is_exclusive(x_63)) { + lean_ctor_release(x_63, 0); + lean_ctor_release(x_63, 1); + x_84 = x_63; } else { - lean_dec_ref(x_61); - x_82 = lean_box(0); + lean_dec_ref(x_63); + 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_56); -x_84 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_84, 0, x_53); -lean_ctor_set(x_84, 1, x_57); -lean_ctor_set(x_84, 2, x_54); -lean_ctor_set(x_84, 3, x_55); -x_1 = x_84; +lean_object* x_86; +lean_dec(x_58); +x_86 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_86, 0, x_55); +lean_ctor_set(x_86, 1, x_59); +lean_ctor_set(x_86, 2, x_56); +lean_ctor_set(x_86, 3, x_57); +x_1 = x_86; goto _start; } } } } } +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_processFacts___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_5); +lean_dec(x_5); +x_13 = l_Lean_Elab_Tactic_Omega_MetaProblem_processFacts(x_1, x_2, x_3, x_4, x_12, x_6, x_7, x_8, x_9, x_10, x_11); +return x_13; +} +} LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_Omega_cases_u2082___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: { @@ -28279,92 +28801,93 @@ lean_dec(x_2); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_Omega_splitDisjunction___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_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_Omega_splitDisjunction___spec__1(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) { _start: { -lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = lean_ctor_get(x_7, 5); -x_11 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_5, x_6, x_7, x_8, x_9); -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) +lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_12 = lean_ctor_get(x_9, 5); +x_13 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_7, x_8, x_9, x_10, x_11); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) { -lean_object* x_13; lean_object* x_14; -x_13 = lean_ctor_get(x_11, 0); -lean_inc(x_10); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_10); -lean_ctor_set(x_14, 1, x_13); -lean_ctor_set_tag(x_11, 1); -lean_ctor_set(x_11, 0, x_14); -return x_11; +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_13, 0); +lean_inc(x_12); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_12); +lean_ctor_set(x_16, 1, x_15); +lean_ctor_set_tag(x_13, 1); +lean_ctor_set(x_13, 0, x_16); +return x_13; } else { -lean_object* x_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); -lean_inc(x_10); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_10); -lean_ctor_set(x_17, 1, x_15); -x_18 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_16); -return x_18; +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_17 = lean_ctor_get(x_13, 0); +x_18 = lean_ctor_get(x_13, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_13); +lean_inc(x_12); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_12); +lean_ctor_set(x_19, 1, x_17); +x_20 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_18); +return x_20; } } } -LEAN_EXPORT lean_object* l_Lean_MVarId_withContext___at_Lean_Elab_Tactic_Omega_splitDisjunction___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, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_MVarId_withContext___at_Lean_Elab_Tactic_Omega_splitDisjunction___spec__2___rarg(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_object* x_11, lean_object* x_12) { _start: { -lean_object* x_11; lean_object* x_12; -x_11 = lean_apply_3(x_2, x_3, x_4, x_5); -x_12 = l___private_Lean_Meta_Basic_0__Lean_Meta_withMVarContextImp___rarg(x_1, x_11, 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_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_box(x_6); +x_14 = lean_apply_5(x_2, x_3, x_4, x_5, x_13, x_7); +x_15 = l___private_Lean_Meta_Basic_0__Lean_Meta_withMVarContextImp___rarg(x_1, x_14, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_15) == 0) { -return x_12; +uint8_t x_16; +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) +{ +return x_15; } 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_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_17; -x_17 = !lean_is_exclusive(x_12); -if (x_17 == 0) +uint8_t x_20; +x_20 = !lean_is_exclusive(x_15); +if (x_20 == 0) { -return x_12; +return x_15; } 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_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; } } } @@ -28373,181 +28896,182 @@ LEAN_EXPORT lean_object* l_Lean_MVarId_withContext___at_Lean_Elab_Tactic_Omega_s _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_MVarId_withContext___at_Lean_Elab_Tactic_Omega_splitDisjunction___spec__2___rarg), 10, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_MVarId_withContext___at_Lean_Elab_Tactic_Omega_splitDisjunction___spec__2___rarg___boxed), 12, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___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_Elab_Tactic_Omega_splitDisjunction___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_object* x_11, lean_object* x_12) { _start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -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_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_13 = lean_ctor_get(x_2, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_2, 1); +lean_inc(x_14); lean_dec(x_2); -x_13 = lean_unsigned_to_nat(0u); -x_14 = lean_nat_dec_lt(x_13, x_12); -lean_dec(x_12); -if (x_14 == 0) +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) { -uint8_t x_15; lean_object* x_16; lean_object* x_17; +uint8_t x_17; lean_object* x_18; lean_object* x_19; +lean_dec(x_13); lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); 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_10); -return x_17; +x_17 = 0; +x_18 = lean_box(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_12); +return x_19; } else { -lean_object* x_18; -x_18 = l_Lean_Elab_Tactic_Omega_omegaImpl(x_11, x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_18) == 0) +lean_object* x_20; +x_20 = l_Lean_Elab_Tactic_Omega_omegaImpl(x_13, x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_20) == 0) { -uint8_t x_19; -x_19 = !lean_is_exclusive(x_18); -if (x_19 == 0) +uint8_t x_21; +x_21 = !lean_is_exclusive(x_20); +if (x_21 == 0) { -lean_object* x_20; uint8_t x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_18, 0); -lean_dec(x_20); -x_21 = 1; -x_22 = lean_box(x_21); -lean_ctor_set(x_18, 0, x_22); -return x_18; +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_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; -x_23 = lean_ctor_get(x_18, 1); -lean_inc(x_23); -lean_dec(x_18); -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; +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_27; -x_27 = !lean_is_exclusive(x_18); -if (x_27 == 0) +uint8_t x_29; +x_29 = !lean_is_exclusive(x_20); +if (x_29 == 0) { -return x_18; +return x_20; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_18, 0); -x_29 = lean_ctor_get(x_18, 1); -lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_18); -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_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; } } } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___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, 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) { _start: { -lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; lean_dec(x_7); -x_16 = lean_st_ref_take(x_12, x_15); -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); -lean_inc(x_18); -lean_dec(x_16); -x_19 = !lean_is_exclusive(x_17); -if (x_19 == 0) +x_18 = lean_st_ref_take(x_14, x_17); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = !lean_is_exclusive(x_19); +if (x_21 == 0) { -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_17, 0); -lean_dec(x_20); -lean_ctor_set(x_17, 0, x_1); -x_21 = lean_st_ref_set(x_12, x_17, x_18); -x_22 = lean_ctor_get(x_21, 1); -lean_inc(x_22); -lean_dec(x_21); -x_23 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_23, 0, x_2); -lean_ctor_set(x_23, 1, x_3); -lean_ctor_set(x_23, 2, x_4); -lean_ctor_set(x_23, 3, x_5); -x_24 = l_Lean_Elab_Tactic_Omega_splitDisjunction(x_23, x_6, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_22); -return x_24; +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_19, 0); +lean_dec(x_22); +lean_ctor_set(x_19, 0, x_1); +x_23 = lean_st_ref_set(x_14, x_19, x_20); +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +lean_dec(x_23); +x_25 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_25, 0, x_2); +lean_ctor_set(x_25, 1, x_3); +lean_ctor_set(x_25, 2, x_4); +lean_ctor_set(x_25, 3, x_5); +x_26 = l_Lean_Elab_Tactic_Omega_splitDisjunction(x_25, x_6, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_24); +return x_26; } 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; -x_25 = lean_ctor_get(x_17, 1); -x_26 = lean_ctor_get(x_17, 2); -x_27 = lean_ctor_get(x_17, 3); +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_27 = lean_ctor_get(x_19, 1); +x_28 = lean_ctor_get(x_19, 2); +x_29 = lean_ctor_get(x_19, 3); +lean_inc(x_29); +lean_inc(x_28); lean_inc(x_27); -lean_inc(x_26); -lean_inc(x_25); -lean_dec(x_17); -x_28 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_28, 0, x_1); -lean_ctor_set(x_28, 1, x_25); -lean_ctor_set(x_28, 2, x_26); -lean_ctor_set(x_28, 3, x_27); -x_29 = lean_st_ref_set(x_12, x_28, x_18); -x_30 = lean_ctor_get(x_29, 1); -lean_inc(x_30); -lean_dec(x_29); -x_31 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_31, 0, x_2); -lean_ctor_set(x_31, 1, x_3); -lean_ctor_set(x_31, 2, x_4); -lean_ctor_set(x_31, 3, x_5); -x_32 = l_Lean_Elab_Tactic_Omega_splitDisjunction(x_31, x_6, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_30); -return x_32; +lean_dec(x_19); +x_30 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_30, 0, x_1); +lean_ctor_set(x_30, 1, x_27); +lean_ctor_set(x_30, 2, x_28); +lean_ctor_set(x_30, 3, x_29); +x_31 = lean_st_ref_set(x_14, x_30, x_20); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +lean_dec(x_31); +x_33 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_33, 0, x_2); +lean_ctor_set(x_33, 1, x_3); +lean_ctor_set(x_33, 2, x_4); +lean_ctor_set(x_33, 3, x_5); +x_34 = l_Lean_Elab_Tactic_Omega_splitDisjunction(x_33, x_6, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_32); +return x_34; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___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, 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) { _start: { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_dec(x_7); -x_16 = l_Lean_Expr_fvar___override(x_1); -x_17 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set(x_17, 1, x_2); -x_18 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_18, 0, x_3); -lean_ctor_set(x_18, 1, x_17); -lean_ctor_set(x_18, 2, x_4); -lean_ctor_set(x_18, 3, x_5); -x_19 = l_Lean_Elab_Tactic_Omega_omegaImpl(x_18, x_6, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -return x_19; +x_18 = l_Lean_Expr_fvar___override(x_1); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_2); +x_20 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_20, 0, x_3); +lean_ctor_set(x_20, 1, x_19); +lean_ctor_set(x_20, 2, x_4); +lean_ctor_set(x_20, 3, x_5); +x_21 = l_Lean_Elab_Tactic_Omega_omegaImpl(x_20, x_6, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +return x_21; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___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_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__4(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) { _start: { -lean_object* x_8; -x_8 = lean_infer_type(x_1, x_3, x_4, x_5, x_6, x_7); -return x_8; +lean_object* x_10; +x_10 = lean_infer_type(x_1, x_5, x_6, x_7, x_8, x_9); +return x_10; } } static lean_object* _init_l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__5___closed__1() { @@ -28584,174 +29108,177 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___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, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___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, 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) { _start: { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_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_12); -x_21 = l_Lean_Expr_fvar___override(x_1); -x_22 = lean_box(0); -x_23 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_23, 0, x_21); -lean_ctor_set(x_23, 1, x_22); +x_23 = l_Lean_Expr_fvar___override(x_1); +x_24 = lean_box(0); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_24 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_24, 0, x_2); -lean_ctor_set(x_24, 1, x_23); -lean_ctor_set(x_24, 2, x_3); -lean_ctor_set(x_24, 3, x_4); -x_25 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_MetaProblem_processFacts), 9, 1); -lean_closure_set(x_25, 0, x_24); +x_26 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_26, 0, x_2); +lean_ctor_set(x_26, 1, x_25); +lean_ctor_set(x_26, 2, x_3); +lean_ctor_set(x_26, 3, x_4); +x_27 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_MetaProblem_processFacts___boxed), 11, 1); +lean_closure_set(x_27, 0, x_26); lean_inc(x_5); -x_26 = lean_alloc_closure((void*)(l_Lean_MVarId_withContext___at_Lean_Elab_Tactic_Omega_splitDisjunction___spec__2___rarg), 10, 2); -lean_closure_set(x_26, 0, x_5); -lean_closure_set(x_26, 1, x_25); -x_27 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__1), 10, 1); -lean_closure_set(x_27, 0, x_5); -x_28 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg), 10, 2); -lean_closure_set(x_28, 0, x_26); +x_28 = lean_alloc_closure((void*)(l_Lean_MVarId_withContext___at_Lean_Elab_Tactic_Omega_splitDisjunction___spec__2___rarg___boxed), 12, 2); +lean_closure_set(x_28, 0, x_5); lean_closure_set(x_28, 1, x_27); +x_29 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__1___boxed), 12, 1); +lean_closure_set(x_29, 0, x_5); +x_30 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg___boxed), 12, 2); +lean_closure_set(x_30, 0, x_28); +lean_closure_set(x_30, 1, x_29); +lean_inc(x_21); +lean_inc(x_20); lean_inc(x_19); lean_inc(x_18); lean_inc(x_17); -lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); -x_29 = l_Lean_Elab_Tactic_Omega_withoutModifyingState___rarg(x_28, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); -if (lean_obj_tag(x_29) == 0) +x_31 = l_Lean_Elab_Tactic_Omega_withoutModifyingState___rarg(x_30, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22); +if (lean_obj_tag(x_31) == 0) { -lean_object* x_30; uint8_t x_31; -x_30 = lean_ctor_get(x_29, 0); -lean_inc(x_30); -x_31 = lean_unbox(x_30); -lean_dec(x_30); -if (x_31 == 0) +lean_object* x_32; uint8_t x_33; +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +x_33 = lean_unbox(x_32); +lean_dec(x_32); +if (x_33 == 0) { -lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; lean_dec(x_11); lean_dec(x_10); -x_32 = lean_ctor_get(x_29, 1); -lean_inc(x_32); -lean_dec(x_29); -lean_inc(x_6); -x_33 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_6, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_32); -x_34 = lean_ctor_get(x_33, 0); +x_34 = lean_ctor_get(x_31, 1); lean_inc(x_34); -x_35 = lean_unbox(x_34); -lean_dec(x_34); -if (x_35 == 0) +lean_dec(x_31); +lean_inc(x_6); +x_35 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_6, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_34); +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +x_37 = lean_unbox(x_36); +lean_dec(x_36); +if (x_37 == 0) { -lean_object* x_36; lean_object* x_37; lean_object* x_38; +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_dec(x_6); -x_36 = lean_ctor_get(x_33, 1); -lean_inc(x_36); -lean_dec(x_33); -x_37 = lean_box(0); -x_38 = l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__2(x_7, x_2, x_8, x_3, x_4, x_9, x_37, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_36); -return x_38; +x_38 = lean_ctor_get(x_35, 1); +lean_inc(x_38); +lean_dec(x_35); +x_39 = lean_box(0); +x_40 = l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__2(x_7, x_2, x_8, x_3, x_4, x_9, x_39, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_38); +return x_40; } else { -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_39 = lean_ctor_get(x_33, 1); -lean_inc(x_39); -lean_dec(x_33); -x_40 = l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__5___closed__2; -x_41 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_6, x_40, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_39); -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 = l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__2(x_7, x_2, x_8, x_3, x_4, x_9, x_42, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_43); -return x_44; +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_41 = lean_ctor_get(x_35, 1); +lean_inc(x_41); +lean_dec(x_35); +x_42 = l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__5___closed__2; +x_43 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_6, x_42, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_41); +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_Elab_Tactic_Omega_splitDisjunction___lambda__2(x_7, x_2, x_8, x_3, x_4, x_9, x_44, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_45); +return x_46; } } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; +lean_object* x_47; lean_object* x_48; lean_object* x_49; uint8_t x_50; lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -x_45 = lean_ctor_get(x_29, 1); -lean_inc(x_45); -lean_dec(x_29); -lean_inc(x_6); -x_46 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_6, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_45); -x_47 = lean_ctor_get(x_46, 0); +x_47 = lean_ctor_get(x_31, 1); lean_inc(x_47); -x_48 = lean_unbox(x_47); -lean_dec(x_47); -if (x_48 == 0) -{ -lean_object* x_49; lean_object* x_50; lean_object* x_51; -lean_dec(x_6); -x_49 = lean_ctor_get(x_46, 1); -lean_inc(x_49); -lean_dec(x_46); -x_50 = lean_box(0); -x_51 = l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__3(x_10, x_22, x_2, x_3, x_4, x_11, x_50, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_49); -return x_51; +lean_dec(x_31); +lean_inc(x_6); +x_48 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_6, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_47); +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_51; lean_object* x_52; lean_object* x_53; +lean_dec(x_6); +x_51 = lean_ctor_get(x_48, 1); +lean_inc(x_51); +lean_dec(x_48); +x_52 = lean_box(0); +x_53 = l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__3(x_10, x_24, x_2, x_3, x_4, x_11, x_52, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_51); +return x_53; } else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_52 = lean_ctor_get(x_46, 1); -lean_inc(x_52); -lean_dec(x_46); +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_54 = lean_ctor_get(x_48, 1); +lean_inc(x_54); +lean_dec(x_48); lean_inc(x_10); -x_53 = l_Lean_Expr_fvar___override(x_10); -x_54 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__4___boxed), 7, 1); -lean_closure_set(x_54, 0, x_53); -x_55 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); -lean_closure_set(x_55, 0, x_54); -x_56 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); +x_55 = l_Lean_Expr_fvar___override(x_10); +x_56 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__4___boxed), 9, 1); lean_closure_set(x_56, 0, x_55); +x_57 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); +lean_closure_set(x_57, 0, x_56); +x_58 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); +lean_closure_set(x_58, 0, x_57); +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_inc(x_13); lean_inc(x_11); -x_57 = l_Lean_MVarId_withContext___at_Lean_Elab_Tactic_Omega_splitDisjunction___spec__2___rarg(x_11, x_56, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_52); -if (lean_obj_tag(x_57) == 0) +x_59 = l_Lean_MVarId_withContext___at_Lean_Elab_Tactic_Omega_splitDisjunction___spec__2___rarg(x_11, x_58, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_54); +if (lean_obj_tag(x_59) == 0) { -lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; -x_58 = lean_ctor_get(x_57, 0); -lean_inc(x_58); -x_59 = lean_ctor_get(x_57, 1); -lean_inc(x_59); -lean_dec(x_57); -x_60 = l_Lean_MessageData_ofExpr(x_58); -x_61 = l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__5___closed__4; -x_62 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_62, 0, x_61); -lean_ctor_set(x_62, 1, x_60); -x_63 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__9; +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_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_MessageData_ofExpr(x_60); +x_63 = l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__5___closed__4; x_64 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_64, 0, x_62); -lean_ctor_set(x_64, 1, x_63); -x_65 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_6, x_64, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_59); -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_Elab_Tactic_Omega_splitDisjunction___lambda__3(x_10, x_22, x_2, x_3, x_4, x_11, x_66, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_67); -return x_68; +lean_ctor_set(x_64, 0, x_63); +lean_ctor_set(x_64, 1, x_62); +x_65 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__9; +x_66 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_66, 0, x_64); +lean_ctor_set(x_66, 1, x_65); +x_67 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_6, x_66, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_61); +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 = l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__3(x_10, x_24, x_2, x_3, x_4, x_11, x_68, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_69); +return x_70; } else { -uint8_t x_69; +uint8_t x_71; +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_15); lean_dec(x_14); lean_dec(x_13); @@ -28761,23 +29288,23 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_69 = !lean_is_exclusive(x_57); -if (x_69 == 0) +x_71 = !lean_is_exclusive(x_59); +if (x_71 == 0) { -return x_57; +return x_59; } else { -lean_object* x_70; lean_object* x_71; lean_object* x_72; -x_70 = lean_ctor_get(x_57, 0); -x_71 = lean_ctor_get(x_57, 1); -lean_inc(x_71); -lean_inc(x_70); -lean_dec(x_57); -x_72 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_72, 0, x_70); -lean_ctor_set(x_72, 1, x_71); -return x_72; +lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_72 = lean_ctor_get(x_59, 0); +x_73 = lean_ctor_get(x_59, 1); +lean_inc(x_73); +lean_inc(x_72); +lean_dec(x_59); +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; } } } @@ -28785,11 +29312,12 @@ return x_72; } else { -uint8_t x_73; +uint8_t x_75; +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_15); lean_dec(x_14); lean_dec(x_13); @@ -28802,23 +29330,23 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_73 = !lean_is_exclusive(x_29); -if (x_73 == 0) +x_75 = !lean_is_exclusive(x_31); +if (x_75 == 0) { -return x_29; +return x_31; } else { -lean_object* x_74; lean_object* x_75; lean_object* x_76; -x_74 = lean_ctor_get(x_29, 0); -x_75 = lean_ctor_get(x_29, 1); -lean_inc(x_75); -lean_inc(x_74); -lean_dec(x_29); -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_object* x_76; lean_object* x_77; lean_object* x_78; +x_76 = lean_ctor_get(x_31, 0); +x_77 = lean_ctor_get(x_31, 1); +lean_inc(x_77); +lean_inc(x_76); +lean_dec(x_31); +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; } } } @@ -28841,128 +29369,130 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___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, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { _start: { -lean_object* x_17; lean_object* x_18; 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_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_dec(x_8); -x_17 = lean_st_ref_get(x_13, 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_ctor_get(x_18, 0); +x_19 = lean_st_ref_get(x_15, x_18); +x_20 = lean_ctor_get(x_19, 0); lean_inc(x_20); -lean_dec(x_18); -x_21 = l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__6___closed__2; +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_Elab_Tactic_Omega_splitDisjunction___lambda__6___closed__2; +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_1); -x_22 = l_Lean_Elab_Tactic_Omega_cases_u2082(x_1, x_2, x_21, x_12, x_13, x_14, x_15, x_19); -if (lean_obj_tag(x_22) == 0) +x_24 = l_Lean_Elab_Tactic_Omega_cases_u2082(x_1, x_2, x_23, x_14, x_15, x_16, x_17, x_21); +if (lean_obj_tag(x_24) == 0) { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; -x_23 = lean_ctor_get(x_22, 0); -lean_inc(x_23); -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_25 = lean_ctor_get(x_24, 0); lean_inc(x_25); -lean_dec(x_23); -x_26 = lean_ctor_get(x_22, 1); +x_26 = lean_ctor_get(x_25, 0); lean_inc(x_26); -lean_dec(x_22); -x_27 = lean_ctor_get(x_24, 0); +x_27 = lean_ctor_get(x_25, 1); lean_inc(x_27); +lean_dec(x_25); x_28 = lean_ctor_get(x_24, 1); lean_inc(x_28); lean_dec(x_24); -x_29 = lean_ctor_get(x_25, 0); +x_29 = lean_ctor_get(x_26, 0); lean_inc(x_29); -x_30 = lean_ctor_get(x_25, 1); +x_30 = lean_ctor_get(x_26, 1); lean_inc(x_30); -lean_dec(x_25); -lean_inc(x_3); -x_31 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_3, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_26); -x_32 = lean_ctor_get(x_31, 0); +lean_dec(x_26); +x_31 = lean_ctor_get(x_27, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_27, 1); lean_inc(x_32); -x_33 = lean_unbox(x_32); -lean_dec(x_32); -if (x_33 == 0) -{ -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_31, 1); +lean_dec(x_27); +lean_inc(x_3); +x_33 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_3, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_28); +x_34 = lean_ctor_get(x_33, 0); lean_inc(x_34); -lean_dec(x_31); -x_35 = lean_box(0); -x_36 = l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__5(x_28, x_4, x_5, x_6, x_27, x_3, x_20, x_7, x_1, x_30, x_29, x_35, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_34); -return x_36; +x_35 = lean_unbox(x_34); +lean_dec(x_34); +if (x_35 == 0) +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_33, 1); +lean_inc(x_36); +lean_dec(x_33); +x_37 = lean_box(0); +x_38 = l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__5(x_30, x_4, x_5, x_6, x_29, x_3, x_22, x_7, x_1, x_32, x_31, x_37, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_36); +return x_38; } 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; -x_37 = lean_ctor_get(x_31, 1); -lean_inc(x_37); -lean_dec(x_31); -lean_inc(x_28); -x_38 = l_Lean_Expr_fvar___override(x_28); -x_39 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__4___boxed), 7, 1); -lean_closure_set(x_39, 0, x_38); -x_40 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); -lean_closure_set(x_40, 0, x_39); -x_41 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 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; +x_39 = lean_ctor_get(x_33, 1); +lean_inc(x_39); +lean_dec(x_33); +lean_inc(x_30); +x_40 = l_Lean_Expr_fvar___override(x_30); +x_41 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__4___boxed), 9, 1); lean_closure_set(x_41, 0, x_40); +x_42 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); +lean_closure_set(x_42, 0, x_41); +x_43 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); +lean_closure_set(x_43, 0, x_42); +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_9); -lean_inc(x_27); -x_42 = l_Lean_MVarId_withContext___at_Lean_Elab_Tactic_Omega_splitDisjunction___spec__2___rarg(x_27, x_41, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_37); -if (lean_obj_tag(x_42) == 0) +lean_inc(x_29); +x_44 = l_Lean_MVarId_withContext___at_Lean_Elab_Tactic_Omega_splitDisjunction___spec__2___rarg(x_29, x_43, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_39); +if (lean_obj_tag(x_44) == 0) { -lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; -x_43 = lean_ctor_get(x_42, 0); -lean_inc(x_43); -x_44 = lean_ctor_get(x_42, 1); -lean_inc(x_44); -lean_dec(x_42); -x_45 = l_Lean_MessageData_ofExpr(x_43); -x_46 = l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__5___closed__4; -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_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__9; +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_45 = lean_ctor_get(x_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_ofExpr(x_45); +x_48 = l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__5___closed__4; x_49 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_49, 0, x_47); -lean_ctor_set(x_49, 1, x_48); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_47); +x_50 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__9; +x_51 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_50); lean_inc(x_3); -x_50 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_3, x_49, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_44); -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 = l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__5(x_28, x_4, x_5, x_6, x_27, x_3, x_20, x_7, x_1, x_30, x_29, x_51, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_52); -return x_53; +x_52 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_3, x_51, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_46); +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_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__5(x_30, x_4, x_5, x_6, x_29, x_3, x_22, x_7, x_1, x_32, x_31, x_53, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_54); +return x_55; } else { -uint8_t x_54; +uint8_t x_56; +lean_dec(x_32); +lean_dec(x_31); lean_dec(x_30); lean_dec(x_29); -lean_dec(x_28); -lean_dec(x_27); -lean_dec(x_20); +lean_dec(x_22); +lean_dec(x_17); +lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); -lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -28972,35 +29502,36 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_54 = !lean_is_exclusive(x_42); -if (x_54 == 0) +x_56 = !lean_is_exclusive(x_44); +if (x_56 == 0) { -return x_42; +return x_44; } else { -lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_55 = lean_ctor_get(x_42, 0); -x_56 = lean_ctor_get(x_42, 1); -lean_inc(x_56); -lean_inc(x_55); -lean_dec(x_42); -x_57 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_57, 0, x_55); -lean_ctor_set(x_57, 1, x_56); -return x_57; +lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_57 = lean_ctor_get(x_44, 0); +x_58 = lean_ctor_get(x_44, 1); +lean_inc(x_58); +lean_inc(x_57); +lean_dec(x_44); +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 { -uint8_t x_58; -lean_dec(x_20); +uint8_t x_60; +lean_dec(x_22); +lean_dec(x_17); +lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); -lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -29010,23 +29541,23 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_58 = !lean_is_exclusive(x_22); -if (x_58 == 0) +x_60 = !lean_is_exclusive(x_24); +if (x_60 == 0) { -return x_22; +return x_24; } else { -lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_59 = lean_ctor_get(x_22, 0); -x_60 = lean_ctor_get(x_22, 1); -lean_inc(x_60); -lean_inc(x_59); -lean_dec(x_22); -x_61 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_61, 0, x_59); -lean_ctor_set(x_61, 1, x_60); -return x_61; +lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_61 = lean_ctor_get(x_24, 0); +x_62 = lean_ctor_get(x_24, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_24); +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; } } } @@ -29048,59 +29579,60 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___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, 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_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___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, uint8_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, uint8_t x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { _start: { if (x_8 == 0) { -lean_object* x_17; lean_object* x_18; -x_17 = lean_box(0); -x_18 = l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_17, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -return x_18; +lean_object* x_19; lean_object* x_20; +x_19 = lean_box(0); +x_20 = l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_19, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +return x_20; } else { -lean_object* x_19; +lean_object* x_21; +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_2); -x_19 = lean_infer_type(x_2, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_19) == 0) +x_21 = lean_infer_type(x_2, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_21) == 0) { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -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_MessageData_ofExpr(x_20); -x_23 = l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__7___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_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__9; -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_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; 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 = 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_MessageData_ofExpr(x_22); +x_25 = l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__7___closed__2; +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_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__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); lean_inc(x_3); -x_27 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_3, x_26, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_21); -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); -lean_inc(x_29); -lean_dec(x_27); -x_30 = l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_28, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_29); -return x_30; +x_29 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_3, x_28, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_23); +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_29, 1); +lean_inc(x_31); +lean_dec(x_29); +x_32 = l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_30, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_31); +return x_32; } else { -uint8_t x_31; +uint8_t x_33; +lean_dec(x_17); +lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); -lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -29111,23 +29643,23 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_31 = !lean_is_exclusive(x_19); -if (x_31 == 0) +x_33 = !lean_is_exclusive(x_21); +if (x_33 == 0) { -return x_19; +return x_21; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_19, 0); -x_33 = lean_ctor_get(x_19, 1); -lean_inc(x_33); -lean_inc(x_32); -lean_dec(x_19); -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_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; } } } @@ -29267,173 +29799,175 @@ static lean_object* _init_l_Lean_Elab_Tactic_Omega_splitDisjunction___closed__14 { lean_object* x_1; lean_object* x_2; x_1 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__2; -x_2 = lean_alloc_closure((void*)(l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3___boxed), 9, 1); +x_2 = lean_alloc_closure((void*)(l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3___boxed), 11, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_Tactic_Omega_splitDisjunction(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_object* x_11, lean_object* x_12) { _start: { -lean_object* x_11; -x_11 = lean_ctor_get(x_1, 2); -lean_inc(x_11); -if (lean_obj_tag(x_11) == 0) +lean_object* x_13; +x_13 = lean_ctor_get(x_1, 2); +lean_inc(x_13); +if (lean_obj_tag(x_13) == 0) { -lean_object* x_12; uint8_t x_13; -x_12 = lean_ctor_get(x_1, 0); -lean_inc(x_12); +lean_object* x_14; uint8_t x_15; +x_14 = lean_ctor_get(x_1, 0); +lean_inc(x_14); lean_dec(x_1); -x_13 = lean_ctor_get_uint8(x_12, sizeof(void*)*7); -if (x_13 == 0) +x_15 = lean_ctor_get_uint8(x_14, sizeof(void*)*7); +if (x_15 == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; -lean_dec(x_12); -x_14 = l_Lean_Elab_Tactic_Omega_splitDisjunction___closed__7; -x_15 = lean_alloc_closure((void*)(l_Lean_throwError___at_Lean_Elab_Tactic_Omega_splitDisjunction___spec__1___boxed), 9, 1); -lean_closure_set(x_15, 0, x_14); -x_16 = l_Lean_MVarId_withContext___at_Lean_Elab_Tactic_Omega_splitDisjunction___spec__2___rarg(x_2, x_15, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_16; +lean_object* x_16; lean_object* x_17; lean_object* x_18; +lean_dec(x_14); +x_16 = l_Lean_Elab_Tactic_Omega_splitDisjunction___closed__7; +x_17 = lean_alloc_closure((void*)(l_Lean_throwError___at_Lean_Elab_Tactic_Omega_splitDisjunction___spec__1___boxed), 11, 1); +lean_closure_set(x_17, 0, x_16); +x_18 = l_Lean_MVarId_withContext___at_Lean_Elab_Tactic_Omega_splitDisjunction___spec__2___rarg(x_2, x_17, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_18; } else { -uint8_t x_17; -x_17 = l_Lean_Elab_Tactic_Omega_Problem_isEmpty(x_12); -if (x_17 == 0) +uint8_t x_19; +x_19 = l_Lean_Elab_Tactic_Omega_Problem_isEmpty(x_14); +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; -x_18 = lean_ctor_get(x_12, 2); -lean_inc(x_18); -lean_dec(x_12); -x_19 = l_Lean_HashMap_toList___at_Lean_Elab_Tactic_Omega_Problem_instToStringProblem___spec__1(x_18); -x_20 = lean_box(0); -x_21 = l_List_mapTR_loop___at_Lean_Elab_Tactic_Omega_Problem_instToStringProblem___spec__4(x_19, x_20); -x_22 = l_Lean_Elab_Tactic_Omega_splitDisjunction___closed__8; -x_23 = l_String_intercalate(x_22, x_21); -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_Elab_Tactic_Omega_splitDisjunction___closed__2; -x_27 = lean_alloc_ctor(7, 2, 0); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_20 = lean_ctor_get(x_14, 2); +lean_inc(x_20); +lean_dec(x_14); +x_21 = l_Lean_HashMap_toList___at_Lean_Elab_Tactic_Omega_Problem_instToStringProblem___spec__1(x_20); +x_22 = lean_box(0); +x_23 = l_List_mapTR_loop___at_Lean_Elab_Tactic_Omega_Problem_instToStringProblem___spec__4(x_21, x_22); +x_24 = l_Lean_Elab_Tactic_Omega_splitDisjunction___closed__8; +x_25 = l_String_intercalate(x_24, x_23); +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_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__9; +x_28 = l_Lean_Elab_Tactic_Omega_splitDisjunction___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_closure((void*)(l_Lean_throwError___at_Lean_Elab_Tactic_Omega_splitDisjunction___spec__1___boxed), 9, 1); -lean_closure_set(x_30, 0, x_29); -x_31 = l_Lean_MVarId_withContext___at_Lean_Elab_Tactic_Omega_splitDisjunction___spec__2___rarg(x_2, x_30, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_31; +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_27); +x_30 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___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 = lean_alloc_closure((void*)(l_Lean_throwError___at_Lean_Elab_Tactic_Omega_splitDisjunction___spec__1___boxed), 11, 1); +lean_closure_set(x_32, 0, x_31); +x_33 = l_Lean_MVarId_withContext___at_Lean_Elab_Tactic_Omega_splitDisjunction___spec__2___rarg(x_2, x_32, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_33; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; -lean_dec(x_12); -x_32 = l_Lean_Elab_Tactic_Omega_splitDisjunction___closed__13; -x_33 = lean_alloc_closure((void*)(l_Lean_throwError___at_Lean_Elab_Tactic_Omega_splitDisjunction___spec__1___boxed), 9, 1); -lean_closure_set(x_33, 0, x_32); -x_34 = l_Lean_MVarId_withContext___at_Lean_Elab_Tactic_Omega_splitDisjunction___spec__2___rarg(x_2, x_33, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_34; +lean_object* x_34; lean_object* x_35; lean_object* x_36; +lean_dec(x_14); +x_34 = l_Lean_Elab_Tactic_Omega_splitDisjunction___closed__13; +x_35 = lean_alloc_closure((void*)(l_Lean_throwError___at_Lean_Elab_Tactic_Omega_splitDisjunction___spec__1___boxed), 11, 1); +lean_closure_set(x_35, 0, x_34); +x_36 = l_Lean_MVarId_withContext___at_Lean_Elab_Tactic_Omega_splitDisjunction___spec__2___rarg(x_2, x_35, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_36; } } } 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; -x_35 = lean_ctor_get(x_1, 0); -lean_inc(x_35); -x_36 = lean_ctor_get(x_1, 1); -lean_inc(x_36); -x_37 = lean_ctor_get(x_1, 3); +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; 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_37 = lean_ctor_get(x_1, 0); lean_inc(x_37); -lean_dec(x_1); -x_38 = lean_ctor_get(x_11, 0); +x_38 = lean_ctor_get(x_1, 1); lean_inc(x_38); -x_39 = lean_ctor_get(x_11, 1); +x_39 = lean_ctor_get(x_1, 3); lean_inc(x_39); -lean_dec(x_11); -x_40 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__2; +lean_dec(x_1); +x_40 = lean_ctor_get(x_13, 0); +lean_inc(x_40); +x_41 = lean_ctor_get(x_13, 1); +lean_inc(x_41); +lean_dec(x_13); +x_42 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__2; lean_inc(x_2); -x_41 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__7___boxed), 16, 7); -lean_closure_set(x_41, 0, x_2); -lean_closure_set(x_41, 1, x_38); -lean_closure_set(x_41, 2, x_40); -lean_closure_set(x_41, 3, x_35); -lean_closure_set(x_41, 4, x_39); -lean_closure_set(x_41, 5, x_37); -lean_closure_set(x_41, 6, x_36); -x_42 = l_Lean_Elab_Tactic_Omega_splitDisjunction___closed__14; -x_43 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg), 10, 2); -lean_closure_set(x_43, 0, x_42); -lean_closure_set(x_43, 1, x_41); -x_44 = l_Lean_MVarId_withContext___at_Lean_Elab_Tactic_Omega_splitDisjunction___spec__2___rarg(x_2, x_43, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_44; +x_43 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__7___boxed), 18, 7); +lean_closure_set(x_43, 0, x_2); +lean_closure_set(x_43, 1, x_40); +lean_closure_set(x_43, 2, x_42); +lean_closure_set(x_43, 3, x_37); +lean_closure_set(x_43, 4, x_41); +lean_closure_set(x_43, 5, x_39); +lean_closure_set(x_43, 6, x_38); +x_44 = l_Lean_Elab_Tactic_Omega_splitDisjunction___closed__14; +x_45 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg___boxed), 12, 2); +lean_closure_set(x_45, 0, x_44); +lean_closure_set(x_45, 1, x_43); +x_46 = l_Lean_MVarId_withContext___at_Lean_Elab_Tactic_Omega_splitDisjunction___spec__2___rarg(x_2, x_45, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_46; } } } -LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at_Lean_Elab_Tactic_Omega_omegaImpl___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at_Lean_Elab_Tactic_Omega_omegaImpl___spec__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_object* x_11, lean_object* x_12) { _start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_11 = lean_st_ref_take(x_7, x_10); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_11, 1); +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_13 = lean_st_ref_take(x_9, x_12); +x_14 = lean_ctor_get(x_13, 0); lean_inc(x_14); -lean_dec(x_11); -x_15 = !lean_is_exclusive(x_12); -if (x_15 == 0) -{ -lean_object* x_16; uint8_t x_17; -x_16 = lean_ctor_get(x_12, 0); -lean_dec(x_16); -x_17 = !lean_is_exclusive(x_13); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +lean_dec(x_13); +x_17 = !lean_is_exclusive(x_14); if (x_17 == 0) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_18 = lean_ctor_get(x_13, 7); -x_19 = l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(x_18, x_1, x_2); -lean_ctor_set(x_13, 7, x_19); -x_20 = lean_st_ref_set(x_7, x_12, x_14); -x_21 = !lean_is_exclusive(x_20); -if (x_21 == 0) +lean_object* x_18; uint8_t x_19; +x_18 = lean_ctor_get(x_14, 0); +lean_dec(x_18); +x_19 = !lean_is_exclusive(x_15); +if (x_19 == 0) { -lean_object* x_22; lean_object* x_23; -x_22 = lean_ctor_get(x_20, 0); -lean_dec(x_22); -x_23 = lean_box(0); -lean_ctor_set(x_20, 0, x_23); -return x_20; +lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_20 = lean_ctor_get(x_15, 7); +x_21 = l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(x_20, x_1, x_2); +lean_ctor_set(x_15, 7, x_21); +x_22 = lean_st_ref_set(x_9, x_14, x_16); +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_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_20, 1); -lean_inc(x_24); -lean_dec(x_20); -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_24); -return x_26; +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_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; 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_13, 0); -x_28 = lean_ctor_get(x_13, 1); -x_29 = lean_ctor_get(x_13, 2); -x_30 = lean_ctor_get(x_13, 3); -x_31 = lean_ctor_get(x_13, 4); -x_32 = lean_ctor_get(x_13, 5); -x_33 = lean_ctor_get(x_13, 6); -x_34 = lean_ctor_get(x_13, 7); -x_35 = lean_ctor_get(x_13, 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; +x_29 = lean_ctor_get(x_15, 0); +x_30 = lean_ctor_get(x_15, 1); +x_31 = lean_ctor_get(x_15, 2); +x_32 = lean_ctor_get(x_15, 3); +x_33 = lean_ctor_get(x_15, 4); +x_34 = lean_ctor_get(x_15, 5); +x_35 = lean_ctor_get(x_15, 6); +x_36 = lean_ctor_get(x_15, 7); +x_37 = lean_ctor_get(x_15, 8); +lean_inc(x_37); +lean_inc(x_36); lean_inc(x_35); lean_inc(x_34); lean_inc(x_33); @@ -29441,143 +29975,142 @@ 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_dec(x_13); -x_36 = l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(x_34, x_1, x_2); -x_37 = lean_alloc_ctor(0, 9, 0); -lean_ctor_set(x_37, 0, x_27); -lean_ctor_set(x_37, 1, x_28); -lean_ctor_set(x_37, 2, x_29); -lean_ctor_set(x_37, 3, x_30); -lean_ctor_set(x_37, 4, x_31); -lean_ctor_set(x_37, 5, x_32); -lean_ctor_set(x_37, 6, x_33); -lean_ctor_set(x_37, 7, x_36); -lean_ctor_set(x_37, 8, x_35); -lean_ctor_set(x_12, 0, x_37); -x_38 = lean_st_ref_set(x_7, x_12, x_14); -x_39 = lean_ctor_get(x_38, 1); -lean_inc(x_39); -if (lean_is_exclusive(x_38)) { - lean_ctor_release(x_38, 0); - lean_ctor_release(x_38, 1); - x_40 = x_38; +lean_dec(x_15); +x_38 = l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(x_36, x_1, x_2); +x_39 = lean_alloc_ctor(0, 9, 0); +lean_ctor_set(x_39, 0, x_29); +lean_ctor_set(x_39, 1, x_30); +lean_ctor_set(x_39, 2, x_31); +lean_ctor_set(x_39, 3, x_32); +lean_ctor_set(x_39, 4, x_33); +lean_ctor_set(x_39, 5, x_34); +lean_ctor_set(x_39, 6, x_35); +lean_ctor_set(x_39, 7, x_38); +lean_ctor_set(x_39, 8, x_37); +lean_ctor_set(x_14, 0, x_39); +x_40 = lean_st_ref_set(x_9, x_14, x_16); +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_38); - x_40 = lean_box(0); + lean_dec_ref(x_40); + x_42 = lean_box(0); } -x_41 = lean_box(0); -if (lean_is_scalar(x_40)) { - x_42 = lean_alloc_ctor(0, 2, 0); +x_43 = lean_box(0); +if (lean_is_scalar(x_42)) { + x_44 = lean_alloc_ctor(0, 2, 0); } else { - x_42 = x_40; + x_44 = x_42; } -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_39); -return x_42; +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_44, 1, x_41); +return x_44; } } else { -lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; -x_43 = lean_ctor_get(x_12, 1); -x_44 = lean_ctor_get(x_12, 2); -x_45 = lean_ctor_get(x_12, 3); -lean_inc(x_45); -lean_inc(x_44); -lean_inc(x_43); -lean_dec(x_12); -x_46 = lean_ctor_get(x_13, 0); -lean_inc(x_46); -x_47 = lean_ctor_get(x_13, 1); +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_45 = lean_ctor_get(x_14, 1); +x_46 = lean_ctor_get(x_14, 2); +x_47 = lean_ctor_get(x_14, 3); lean_inc(x_47); -x_48 = lean_ctor_get(x_13, 2); +lean_inc(x_46); +lean_inc(x_45); +lean_dec(x_14); +x_48 = lean_ctor_get(x_15, 0); lean_inc(x_48); -x_49 = lean_ctor_get(x_13, 3); +x_49 = lean_ctor_get(x_15, 1); lean_inc(x_49); -x_50 = lean_ctor_get(x_13, 4); +x_50 = lean_ctor_get(x_15, 2); lean_inc(x_50); -x_51 = lean_ctor_get(x_13, 5); +x_51 = lean_ctor_get(x_15, 3); lean_inc(x_51); -x_52 = lean_ctor_get(x_13, 6); +x_52 = lean_ctor_get(x_15, 4); lean_inc(x_52); -x_53 = lean_ctor_get(x_13, 7); +x_53 = lean_ctor_get(x_15, 5); lean_inc(x_53); -x_54 = lean_ctor_get(x_13, 8); +x_54 = lean_ctor_get(x_15, 6); lean_inc(x_54); -if (lean_is_exclusive(x_13)) { - lean_ctor_release(x_13, 0); - lean_ctor_release(x_13, 1); - lean_ctor_release(x_13, 2); - lean_ctor_release(x_13, 3); - lean_ctor_release(x_13, 4); - lean_ctor_release(x_13, 5); - lean_ctor_release(x_13, 6); - lean_ctor_release(x_13, 7); - lean_ctor_release(x_13, 8); - x_55 = x_13; +x_55 = lean_ctor_get(x_15, 7); +lean_inc(x_55); +x_56 = lean_ctor_get(x_15, 8); +lean_inc(x_56); +if (lean_is_exclusive(x_15)) { + lean_ctor_release(x_15, 0); + lean_ctor_release(x_15, 1); + lean_ctor_release(x_15, 2); + lean_ctor_release(x_15, 3); + lean_ctor_release(x_15, 4); + lean_ctor_release(x_15, 5); + lean_ctor_release(x_15, 6); + lean_ctor_release(x_15, 7); + lean_ctor_release(x_15, 8); + x_57 = x_15; } else { - lean_dec_ref(x_13); - x_55 = lean_box(0); + lean_dec_ref(x_15); + x_57 = lean_box(0); } -x_56 = l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(x_53, x_1, x_2); -if (lean_is_scalar(x_55)) { - x_57 = lean_alloc_ctor(0, 9, 0); +x_58 = l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(x_55, x_1, x_2); +if (lean_is_scalar(x_57)) { + x_59 = lean_alloc_ctor(0, 9, 0); } else { - x_57 = x_55; -} -lean_ctor_set(x_57, 0, x_46); -lean_ctor_set(x_57, 1, x_47); -lean_ctor_set(x_57, 2, x_48); -lean_ctor_set(x_57, 3, x_49); -lean_ctor_set(x_57, 4, x_50); -lean_ctor_set(x_57, 5, x_51); -lean_ctor_set(x_57, 6, x_52); -lean_ctor_set(x_57, 7, x_56); -lean_ctor_set(x_57, 8, x_54); -x_58 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_58, 0, x_57); -lean_ctor_set(x_58, 1, x_43); -lean_ctor_set(x_58, 2, x_44); -lean_ctor_set(x_58, 3, x_45); -x_59 = lean_st_ref_set(x_7, x_58, x_14); -x_60 = lean_ctor_get(x_59, 1); -lean_inc(x_60); -if (lean_is_exclusive(x_59)) { - lean_ctor_release(x_59, 0); - lean_ctor_release(x_59, 1); - x_61 = x_59; + x_59 = x_57; +} +lean_ctor_set(x_59, 0, x_48); +lean_ctor_set(x_59, 1, x_49); +lean_ctor_set(x_59, 2, x_50); +lean_ctor_set(x_59, 3, x_51); +lean_ctor_set(x_59, 4, x_52); +lean_ctor_set(x_59, 5, x_53); +lean_ctor_set(x_59, 6, x_54); +lean_ctor_set(x_59, 7, x_58); +lean_ctor_set(x_59, 8, x_56); +x_60 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_60, 0, x_59); +lean_ctor_set(x_60, 1, x_45); +lean_ctor_set(x_60, 2, x_46); +lean_ctor_set(x_60, 3, x_47); +x_61 = lean_st_ref_set(x_9, x_60, x_16); +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_59); - x_61 = lean_box(0); + lean_dec_ref(x_61); + x_63 = lean_box(0); } -x_62 = lean_box(0); -if (lean_is_scalar(x_61)) { - x_63 = lean_alloc_ctor(0, 2, 0); +x_64 = lean_box(0); +if (lean_is_scalar(x_63)) { + x_65 = lean_alloc_ctor(0, 2, 0); } else { - x_63 = x_61; + x_65 = x_63; } -lean_ctor_set(x_63, 0, x_62); -lean_ctor_set(x_63, 1, x_60); -return 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_Lean_Elab_Tactic_Omega_omegaImpl___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_Lean_Elab_Tactic_Omega_omegaImpl___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, 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) { _start: { -lean_object* x_12; +lean_object* x_14; lean_dec(x_3); -x_12 = l_Lean_MVarId_assign___at_Lean_Elab_Tactic_Omega_omegaImpl___spec__1(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_14 = l_Lean_MVarId_assign___at_Lean_Elab_Tactic_Omega_omegaImpl___spec__1(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -return x_12; +return x_14; } } static lean_object* _init_l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__2___closed__1() { @@ -29597,162 +30130,168 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___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, lean_object* x_13, lean_object* x_14) { _start: { -lean_object* x_13; +lean_object* x_15; lean_object* x_16; lean_dec(x_4); +x_15 = lean_box(x_8); +lean_inc(x_13); +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); -lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_13 = lean_apply_8(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_13) == 0) +x_16 = lean_apply_10(x_1, x_5, x_6, x_7, x_15, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_16) == 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; uint8_t x_21; -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); -x_16 = l_Lean_instantiateMVars___at_Lean_Elab_Tactic_Omega_MetaProblem_addFact___spec__1(x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_15); +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; x_17 = lean_ctor_get(x_16, 0); lean_inc(x_17); x_18 = lean_ctor_get(x_16, 1); lean_inc(x_18); lean_dec(x_16); -lean_inc(x_2); -x_19 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_18); +x_19 = l_Lean_instantiateMVars___at_Lean_Elab_Tactic_Omega_MetaProblem_addFact___spec__1(x_17, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_18); x_20 = lean_ctor_get(x_19, 0); lean_inc(x_20); -x_21 = lean_unbox(x_20); -lean_dec(x_20); -if (x_21 == 0) -{ -lean_object* x_22; lean_object* x_23; -lean_dec(x_2); -x_22 = lean_ctor_get(x_19, 1); -lean_inc(x_22); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); lean_dec(x_19); -x_23 = l_Lean_MVarId_assign___at_Lean_Elab_Tactic_Omega_omegaImpl___spec__1(x_3, x_17, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_22); +lean_inc(x_2); +x_22 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_21); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_unbox(x_23); +lean_dec(x_23); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; +lean_dec(x_2); +x_25 = lean_ctor_get(x_22, 1); +lean_inc(x_25); +lean_dec(x_22); +x_26 = l_Lean_MVarId_assign___at_Lean_Elab_Tactic_Omega_omegaImpl___spec__1(x_3, x_20, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_25); +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); -return x_23; +return x_26; } else { -lean_object* x_24; lean_object* x_25; -x_24 = lean_ctor_get(x_19, 1); -lean_inc(x_24); -lean_dec(x_19); +lean_object* x_27; lean_object* x_28; +x_27 = lean_ctor_get(x_22, 1); +lean_inc(x_27); +lean_dec(x_22); +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_17); -x_25 = lean_infer_type(x_17, x_8, x_9, x_10, x_11, x_24); -if (lean_obj_tag(x_25) == 0) +lean_inc(x_20); +x_28 = lean_infer_type(x_20, x_10, x_11, x_12, x_13, x_27); +if (lean_obj_tag(x_28) == 0) { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_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_MessageData_ofExpr(x_26); -x_29 = l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__2___closed__2; -x_30 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_30, 0, x_29); -lean_ctor_set(x_30, 1, x_28); -x_31 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__9; -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 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_2, x_32, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_27); -x_34 = lean_ctor_get(x_33, 1); -lean_inc(x_34); -lean_dec(x_33); -x_35 = l_Lean_MVarId_assign___at_Lean_Elab_Tactic_Omega_omegaImpl___spec__1(x_3, x_17, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_34); +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_29 = lean_ctor_get(x_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_Lean_MessageData_ofExpr(x_29); +x_32 = l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__2___closed__2; +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_Omega_asLinearComboImpl_rewrite___closed__9; +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_Tactic_Omega_lookup___spec__14(x_2, x_35, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_30); +x_37 = lean_ctor_get(x_36, 1); +lean_inc(x_37); +lean_dec(x_36); +x_38 = l_Lean_MVarId_assign___at_Lean_Elab_Tactic_Omega_omegaImpl___spec__1(x_3, x_20, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_37); +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); -return x_35; +return x_38; } else { -uint8_t x_36; -lean_dec(x_17); +uint8_t x_39; +lean_dec(x_20); +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); -x_36 = !lean_is_exclusive(x_25); -if (x_36 == 0) +x_39 = !lean_is_exclusive(x_28); +if (x_39 == 0) { -return x_25; +return x_28; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_25, 0); -x_38 = lean_ctor_get(x_25, 1); -lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_25); -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_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_28, 0); +x_41 = lean_ctor_get(x_28, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_28); +x_42 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_41); +return x_42; } } } } else { -uint8_t x_40; +uint8_t x_43; +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); -x_40 = !lean_is_exclusive(x_13); -if (x_40 == 0) +x_43 = !lean_is_exclusive(x_16); +if (x_43 == 0) { -return x_13; +return x_16; } else { -lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_41 = lean_ctor_get(x_13, 0); -x_42 = lean_ctor_get(x_13, 1); -lean_inc(x_42); -lean_inc(x_41); -lean_dec(x_13); -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_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_16, 0); +x_45 = lean_ctor_get(x_16, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_16); +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; } } } @@ -29774,77 +30313,77 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___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_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___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, 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) { _start: { -uint8_t x_14; +uint8_t x_16; lean_dec(x_5); -x_14 = lean_ctor_get_uint8(x_1, sizeof(void*)*7); -if (x_14 == 0) +x_16 = lean_ctor_get_uint8(x_1, sizeof(void*)*7); +if (x_16 == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; lean_dec(x_4); -x_15 = lean_ctor_get(x_1, 5); -lean_inc(x_15); -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -lean_dec(x_15); -lean_inc(x_2); -x_17 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +x_17 = lean_ctor_get(x_1, 5); +lean_inc(x_17); 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_dec(x_17); +lean_inc(x_2); +x_19 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_unbox(x_20); +lean_dec(x_20); +if (x_21 == 0) { -lean_object* x_20; lean_object* x_21; lean_object* x_22; +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_dec(x_1); -x_20 = lean_ctor_get(x_17, 1); -lean_inc(x_20); -lean_dec(x_17); -x_21 = lean_box(0); -x_22 = l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__2(x_16, x_2, x_3, x_21, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_20); -return x_22; +x_22 = lean_ctor_get(x_19, 1); +lean_inc(x_22); +lean_dec(x_19); +x_23 = lean_box(0); +x_24 = l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__2(x_18, x_2, x_3, x_23, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_22); +return x_24; } 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_ctor_get(x_17, 1); -lean_inc(x_23); -lean_dec(x_17); -x_24 = lean_ctor_get(x_1, 6); -lean_inc(x_24); +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; 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_19, 1); +lean_inc(x_25); +lean_dec(x_19); +x_26 = lean_ctor_get(x_1, 6); +lean_inc(x_26); lean_dec(x_1); -x_25 = lean_thunk_get_own(x_24); -lean_dec(x_24); -x_26 = l_Lean_stringToMessageData(x_25); -lean_dec(x_25); -x_27 = l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__3___closed__2; -x_28 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_28, 0, x_27); -lean_ctor_set(x_28, 1, x_26); -x_29 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__9; +x_27 = lean_thunk_get_own(x_26); +lean_dec(x_26); +x_28 = l_Lean_stringToMessageData(x_27); +lean_dec(x_27); +x_29 = l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__3___closed__2; x_30 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_30, 0, x_28); -lean_ctor_set(x_30, 1, x_29); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_28); +x_31 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__9; +x_32 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); lean_inc(x_2); -x_31 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_2, x_30, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_23); -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_Lean_Elab_Tactic_Omega_omegaImpl___lambda__2(x_16, x_2, x_3, x_32, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_33); -return x_34; +x_33 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_2, x_32, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_25); +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_Lean_Elab_Tactic_Omega_omegaImpl___lambda__2(x_18, x_2, x_3, x_34, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_35); +return x_36; } } else { -lean_object* x_35; +lean_object* x_37; lean_dec(x_2); lean_dec(x_1); -x_35 = l_Lean_Elab_Tactic_Omega_splitDisjunction(x_4, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_35; +x_37 = l_Lean_Elab_Tactic_Omega_splitDisjunction(x_4, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_37; } } } @@ -29874,141 +30413,142 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___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_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___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) { _start: { -lean_object* x_13; lean_object* x_14; uint8_t x_15; +lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_inc(x_1); -x_13 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_unbox(x_14); -lean_dec(x_14); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_13, 1); +x_15 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +x_16 = lean_ctor_get(x_15, 0); lean_inc(x_16); -lean_dec(x_13); -x_17 = lean_box(0); -x_18 = l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__3(x_4, x_1, x_2, x_3, x_17, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_16); -return x_18; +x_17 = lean_unbox(x_16); +lean_dec(x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_15, 1); +lean_inc(x_18); +lean_dec(x_15); +x_19 = lean_box(0); +x_20 = l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__3(x_4, 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_18); +return x_20; } 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_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_15, 1); +lean_inc(x_21); +lean_dec(x_15); +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_20 = l_Lean_Elab_Tactic_Omega_atomsList___rarg(x_6, x_7, x_8, x_9, x_10, x_11, x_19); -if (lean_obj_tag(x_20) == 0) +x_22 = l_Lean_Elab_Tactic_Omega_atomsList___rarg(x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, 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; uint8_t x_28; -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_MessageData_ofExpr(x_21); -x_24 = l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__4___closed__2; -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_Tactic_Omega_omegaImpl___lambda__4___closed__3; +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; +x_23 = lean_ctor_get(x_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_MessageData_ofExpr(x_23); +x_26 = l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__4___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 = lean_ctor_get_uint8(x_4, sizeof(void*)*7); -if (x_28 == 0) +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_25); +x_28 = l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__4___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 = lean_ctor_get_uint8(x_4, sizeof(void*)*7); +if (x_30 == 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; -x_29 = l_Lean_Elab_Tactic_Omega_splitDisjunction___closed__5; -x_30 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_30, 0, x_27); -lean_ctor_set(x_30, 1, x_29); -x_31 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__9; +lean_object* x_31; lean_object* x_32; 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_31 = l_Lean_Elab_Tactic_Omega_splitDisjunction___closed__5; x_32 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 0, x_29); lean_ctor_set(x_32, 1, x_31); +x_33 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__9; +x_34 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); lean_inc(x_1); -x_33 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_1, x_32, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_22); -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_Lean_Elab_Tactic_Omega_omegaImpl___lambda__3(x_4, x_1, x_2, x_3, x_34, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_35); -return x_36; +x_35 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_1, x_34, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_24); +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_Elab_Tactic_Omega_omegaImpl___lambda__3(x_4, x_1, x_2, x_3, x_36, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_37); +return x_38; } else { -uint8_t x_37; -x_37 = l_Lean_Elab_Tactic_Omega_Problem_isEmpty(x_4); -if (x_37 == 0) +uint8_t x_39; +x_39 = l_Lean_Elab_Tactic_Omega_Problem_isEmpty(x_4); +if (x_39 == 0) { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_38 = lean_ctor_get(x_4, 2); -lean_inc(x_38); -x_39 = l_Lean_HashMap_toList___at_Lean_Elab_Tactic_Omega_Problem_instToStringProblem___spec__1(x_38); -x_40 = lean_box(0); -x_41 = l_List_mapTR_loop___at_Lean_Elab_Tactic_Omega_Problem_instToStringProblem___spec__4(x_39, x_40); -x_42 = l_Lean_Elab_Tactic_Omega_splitDisjunction___closed__8; -x_43 = l_String_intercalate(x_42, x_41); -x_44 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_44, 0, x_43); -x_45 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_45, 0, x_44); -x_46 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_46, 0, x_27); -lean_ctor_set(x_46, 1, x_45); -x_47 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__9; +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_40 = lean_ctor_get(x_4, 2); +lean_inc(x_40); +x_41 = l_Lean_HashMap_toList___at_Lean_Elab_Tactic_Omega_Problem_instToStringProblem___spec__1(x_40); +x_42 = lean_box(0); +x_43 = l_List_mapTR_loop___at_Lean_Elab_Tactic_Omega_Problem_instToStringProblem___spec__4(x_41, x_42); +x_44 = l_Lean_Elab_Tactic_Omega_splitDisjunction___closed__8; +x_45 = l_String_intercalate(x_44, x_43); +x_46 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_46, 0, x_45); +x_47 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_47, 0, x_46); x_48 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_48, 0, x_46); +lean_ctor_set(x_48, 0, x_29); lean_ctor_set(x_48, 1, x_47); +x_49 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__9; +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_49 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_1, x_48, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_22); -x_50 = lean_ctor_get(x_49, 0); -lean_inc(x_50); -x_51 = lean_ctor_get(x_49, 1); -lean_inc(x_51); -lean_dec(x_49); -x_52 = l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__3(x_4, x_1, x_2, x_3, x_50, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_51); -return x_52; +x_51 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_1, x_50, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_24); +x_52 = lean_ctor_get(x_51, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_51, 1); +lean_inc(x_53); +lean_dec(x_51); +x_54 = l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__3(x_4, x_1, x_2, x_3, x_52, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_53); +return x_54; } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_53 = l_Lean_Elab_Tactic_Omega_splitDisjunction___closed__11; -x_54 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_54, 0, x_27); -lean_ctor_set(x_54, 1, x_53); -x_55 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__9; +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_Elab_Tactic_Omega_splitDisjunction___closed__11; x_56 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_56, 0, x_54); +lean_ctor_set(x_56, 0, x_29); lean_ctor_set(x_56, 1, x_55); +x_57 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__9; +x_58 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_58, 0, x_56); +lean_ctor_set(x_58, 1, x_57); lean_inc(x_1); -x_57 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_1, x_56, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_22); -x_58 = lean_ctor_get(x_57, 0); -lean_inc(x_58); -x_59 = lean_ctor_get(x_57, 1); -lean_inc(x_59); -lean_dec(x_57); -x_60 = l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__3(x_4, x_1, x_2, x_3, x_58, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_59); -return x_60; +x_59 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_1, x_58, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_24); +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_Elab_Tactic_Omega_omegaImpl___lambda__3(x_4, x_1, x_2, x_3, x_60, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_61); +return x_62; } } } else { -uint8_t x_61; +uint8_t x_63; +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); @@ -30016,92 +30556,94 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_61 = !lean_is_exclusive(x_20); -if (x_61 == 0) +x_63 = !lean_is_exclusive(x_22); +if (x_63 == 0) { -return x_20; +return x_22; } else { -lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_62 = lean_ctor_get(x_20, 0); -x_63 = lean_ctor_get(x_20, 1); -lean_inc(x_63); -lean_inc(x_62); -lean_dec(x_20); -x_64 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_64, 0, x_62); -lean_ctor_set(x_64, 1, x_63); -return x_64; +lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_64 = lean_ctor_get(x_22, 0); +x_65 = lean_ctor_get(x_22, 1); +lean_inc(x_65); +lean_inc(x_64); +lean_dec(x_22); +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_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___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_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___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, 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) { _start: { -uint8_t x_14; +uint8_t x_16; lean_dec(x_5); -x_14 = lean_ctor_get_uint8(x_4, sizeof(void*)*7); -if (x_14 == 0) +x_16 = lean_ctor_get_uint8(x_4, sizeof(void*)*7); +if (x_16 == 0) { -lean_object* x_15; -x_15 = l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__4(x_1, x_2, x_3, x_4, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_15; +lean_object* x_17; +x_17 = l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__4(x_1, x_2, x_3, x_4, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_17; } else { -lean_object* x_16; +lean_object* x_18; +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); lean_inc(x_7); lean_inc(x_6); -x_16 = l_Lean_Elab_Tactic_Omega_Problem_elimination(x_4, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_16) == 0) +x_18 = l_Lean_Elab_Tactic_Omega_Problem_elimination(x_4, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); -lean_inc(x_18); -lean_dec(x_16); -x_19 = l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__4(x_1, x_2, x_3, x_17, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_18); -return x_19; +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__4(x_1, x_2, x_3, x_19, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_20); +return x_21; } else { -uint8_t x_20; +uint8_t x_22; +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_3); lean_dec(x_2); lean_dec(x_1); -x_20 = !lean_is_exclusive(x_16); -if (x_20 == 0) +x_22 = !lean_is_exclusive(x_18); +if (x_22 == 0) { -return x_16; +return x_18; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_16, 0); -x_22 = lean_ctor_get(x_16, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_16); -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_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_18, 0); +x_24 = lean_ctor_get(x_18, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_18); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; } } } @@ -30141,250 +30683,332 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___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_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__6(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_object* x_11, lean_object* x_12) { _start: { -lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_11 = lean_ctor_get(x_2, 0); -lean_inc(x_11); +lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_13 = lean_ctor_get(x_2, 0); +lean_inc(x_13); lean_dec(x_2); -x_12 = lean_ctor_get(x_11, 1); -lean_inc(x_12); -x_13 = l_List_isEmpty___rarg(x_12); -lean_dec(x_12); -if (x_13 == 0) +x_14 = lean_ctor_get(x_13, 1); +lean_inc(x_14); +x_15 = l_List_isEmpty___rarg(x_14); +lean_dec(x_14); +if (x_15 == 0) { -lean_object* x_14; lean_object* x_15; uint8_t x_16; -lean_dec(x_11); +lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_dec(x_13); +lean_dec(x_7); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_14 = l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__6___closed__2; -x_15 = l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(x_14, x_6, x_7, x_8, x_9, x_10); +x_16 = l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__6___closed__2; +x_17 = l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(x_16, 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); -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) { -return x_15; +return x_17; } 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(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_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(1, 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_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_20 = lean_ctor_get(x_11, 0); -lean_inc(x_20); -x_21 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__2; -x_22 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_21, x_3, 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_unbox(x_23); -lean_dec(x_23); -if (x_24 == 0) -{ -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_22, 1); +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_13, 0); +lean_inc(x_22); +x_23 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__2; +x_24 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_23, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_25 = lean_ctor_get(x_24, 0); lean_inc(x_25); -lean_dec(x_22); -x_26 = lean_box(0); -x_27 = l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__5(x_21, x_1, x_11, x_20, x_26, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_25); -return x_27; +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; +x_27 = lean_ctor_get(x_24, 1); +lean_inc(x_27); +lean_dec(x_24); +x_28 = lean_box(0); +x_29 = l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__5(x_23, x_1, x_13, x_22, x_28, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_27); +return x_29; } else { -lean_object* x_28; lean_object* x_29; -x_28 = lean_ctor_get(x_22, 1); -lean_inc(x_28); -lean_dec(x_22); +lean_object* x_30; lean_object* x_31; +x_30 = lean_ctor_get(x_24, 1); +lean_inc(x_30); +lean_dec(x_24); +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_29 = l_Lean_Elab_Tactic_Omega_atomsList___rarg(x_4, x_5, x_6, x_7, x_8, x_9, x_28); -if (lean_obj_tag(x_29) == 0) +x_31 = l_Lean_Elab_Tactic_Omega_atomsList___rarg(x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_30); +if (lean_obj_tag(x_31) == 0) { -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_30 = lean_ctor_get(x_29, 0); -lean_inc(x_30); -x_31 = lean_ctor_get(x_29, 1); -lean_inc(x_31); -lean_dec(x_29); -x_32 = l_Lean_MessageData_ofExpr(x_30); -x_33 = l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__6___closed__4; -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_Elab_Tactic_Omega_omegaImpl___lambda__4___closed__3; +lean_object* x_32; 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_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_Lean_MessageData_ofExpr(x_32); +x_35 = l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__6___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 = lean_ctor_get_uint8(x_20, sizeof(void*)*7); -if (x_37 == 0) +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_34); +x_37 = l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__4___closed__3; +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_uint8(x_22, sizeof(void*)*7); +if (x_39 == 0) { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_38 = l_Lean_Elab_Tactic_Omega_splitDisjunction___closed__5; -x_39 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__9; +lean_object* x_40; lean_object* x_41; 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 = l_Lean_Elab_Tactic_Omega_splitDisjunction___closed__5; x_41 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 0, x_38); lean_ctor_set(x_41, 1, x_40); -x_42 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_21, x_41, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_31); -x_43 = lean_ctor_get(x_42, 0); -lean_inc(x_43); -x_44 = lean_ctor_get(x_42, 1); -lean_inc(x_44); -lean_dec(x_42); -x_45 = l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__5(x_21, x_1, x_11, x_20, x_43, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_44); -return x_45; +x_42 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__9; +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_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_23, x_43, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_33); +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_Elab_Tactic_Omega_omegaImpl___lambda__5(x_23, x_1, x_13, x_22, x_45, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_46); +return x_47; } else { -uint8_t x_46; -x_46 = l_Lean_Elab_Tactic_Omega_Problem_isEmpty(x_20); -if (x_46 == 0) +uint8_t x_48; +x_48 = l_Lean_Elab_Tactic_Omega_Problem_isEmpty(x_22); +if (x_48 == 0) { -lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_47 = lean_ctor_get(x_20, 2); -lean_inc(x_47); -x_48 = l_Lean_HashMap_toList___at_Lean_Elab_Tactic_Omega_Problem_instToStringProblem___spec__1(x_47); -x_49 = lean_box(0); -x_50 = l_List_mapTR_loop___at_Lean_Elab_Tactic_Omega_Problem_instToStringProblem___spec__4(x_48, x_49); -x_51 = l_Lean_Elab_Tactic_Omega_splitDisjunction___closed__8; -x_52 = l_String_intercalate(x_51, x_50); -x_53 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_53, 0, x_52); -x_54 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_54, 0, x_53); -x_55 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_55, 0, x_36); -lean_ctor_set(x_55, 1, x_54); -x_56 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__9; +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_49 = lean_ctor_get(x_22, 2); +lean_inc(x_49); +x_50 = l_Lean_HashMap_toList___at_Lean_Elab_Tactic_Omega_Problem_instToStringProblem___spec__1(x_49); +x_51 = lean_box(0); +x_52 = l_List_mapTR_loop___at_Lean_Elab_Tactic_Omega_Problem_instToStringProblem___spec__4(x_50, x_51); +x_53 = l_Lean_Elab_Tactic_Omega_splitDisjunction___closed__8; +x_54 = l_String_intercalate(x_53, x_52); +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_55); +lean_ctor_set(x_57, 0, x_38); lean_ctor_set(x_57, 1, x_56); -x_58 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_21, x_57, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_31); -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_Tactic_Omega_omegaImpl___lambda__5(x_21, x_1, x_11, x_20, x_59, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_60); -return x_61; +x_58 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___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_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_23, x_59, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_33); +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_Elab_Tactic_Omega_omegaImpl___lambda__5(x_23, x_1, x_13, x_22, x_61, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_62); +return x_63; } else { -lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; -x_62 = l_Lean_Elab_Tactic_Omega_splitDisjunction___closed__11; -x_63 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_63, 0, x_36); -lean_ctor_set(x_63, 1, x_62); -x_64 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__9; +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_64 = l_Lean_Elab_Tactic_Omega_splitDisjunction___closed__11; x_65 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_65, 0, x_63); +lean_ctor_set(x_65, 0, x_38); lean_ctor_set(x_65, 1, x_64); -x_66 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_21, x_65, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_31); -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_Elab_Tactic_Omega_omegaImpl___lambda__5(x_21, x_1, x_11, x_20, x_67, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_68); -return x_69; +x_66 = l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__9; +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_Elab_Tactic_Omega_lookup___spec__14(x_23, x_67, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_33); +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_Elab_Tactic_Omega_omegaImpl___lambda__5(x_23, x_1, x_13, x_22, x_69, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_70); +return x_71; } } } else { -uint8_t x_70; -lean_dec(x_20); +uint8_t x_72; +lean_dec(x_22); +lean_dec(x_13); lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_70 = !lean_is_exclusive(x_29); -if (x_70 == 0) +x_72 = !lean_is_exclusive(x_31); +if (x_72 == 0) { -return x_29; +return x_31; } else { -lean_object* x_71; lean_object* x_72; lean_object* x_73; -x_71 = lean_ctor_get(x_29, 0); -x_72 = lean_ctor_get(x_29, 1); -lean_inc(x_72); -lean_inc(x_71); -lean_dec(x_29); -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_73; lean_object* x_74; lean_object* x_75; +x_73 = lean_ctor_get(x_31, 0); +x_74 = lean_ctor_get(x_31, 1); +lean_inc(x_74); +lean_inc(x_73); +lean_dec(x_31); +x_75 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_75, 0, x_73); +lean_ctor_set(x_75, 1, x_74); +return x_75; } } } } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_Tactic_Omega_omegaImpl(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_object* x_11, lean_object* x_12) { _start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_11 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_MetaProblem_processFacts), 9, 1); -lean_closure_set(x_11, 0, x_1); +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_13 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_MetaProblem_processFacts___boxed), 11, 1); +lean_closure_set(x_13, 0, x_1); lean_inc(x_2); -x_12 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__6), 10, 1); -lean_closure_set(x_12, 0, x_2); -x_13 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg), 10, 2); -lean_closure_set(x_13, 0, x_11); -lean_closure_set(x_13, 1, x_12); -x_14 = l_Lean_MVarId_withContext___at_Lean_Elab_Tactic_Omega_splitDisjunction___spec__2___rarg(x_2, x_13, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_14; +x_14 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__6___boxed), 12, 1); +lean_closure_set(x_14, 0, x_2); +x_15 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg___boxed), 12, 2); +lean_closure_set(x_15, 0, x_13); +lean_closure_set(x_15, 1, x_14); +x_16 = l_Lean_MVarId_withContext___at_Lean_Elab_Tactic_Omega_splitDisjunction___spec__2___rarg(x_2, x_15, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_16; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_Omega_splitDisjunction___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_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_Omega_splitDisjunction___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_10; -x_10 = l_Lean_throwError___at_Lean_Elab_Tactic_Omega_splitDisjunction___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +uint8_t x_12; lean_object* x_13; +x_12 = lean_unbox(x_5); +lean_dec(x_5); +x_13 = l_Lean_throwError___at_Lean_Elab_Tactic_Omega_splitDisjunction___spec__1(x_1, x_2, x_3, x_4, x_12, 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_10; +return x_13; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___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_MVarId_withContext___at_Lean_Elab_Tactic_Omega_splitDisjunction___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, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_8; -x_8 = l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +uint8_t x_13; lean_object* x_14; +x_13 = lean_unbox(x_6); +lean_dec(x_6); +x_14 = l_Lean_MVarId_withContext___at_Lean_Elab_Tactic_Omega_splitDisjunction___spec__2___rarg(x_1, x_2, x_3, x_4, x_5, x_13, x_7, x_8, x_9, x_10, x_11, x_12); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; lean_object* x_14; +x_13 = lean_unbox(x_6); +lean_dec(x_6); +x_14 = l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__1(x_1, x_2, x_3, x_4, x_5, x_13, x_7, x_8, x_9, x_10, x_11, x_12); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__2___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +_start: +{ +uint8_t x_18; lean_object* x_19; +x_18 = lean_unbox(x_11); +lean_dec(x_11); +x_19 = l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_18, x_12, x_13, x_14, x_15, x_16, x_17); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___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]; +_start: +{ +uint8_t x_18; lean_object* x_19; +x_18 = lean_unbox(x_11); +lean_dec(x_11); +x_19 = l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_18, x_12, x_13, x_14, x_15, x_16, x_17); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +uint8_t x_10; lean_object* x_11; +x_10 = lean_unbox(x_3); +lean_dec(x_3); +x_11 = l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__4(x_1, x_2, x_10, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_4); lean_dec(x_2); -return x_8; +return x_11; } } LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__5___boxed(lean_object** _args) { @@ -30408,36 +31032,171 @@ lean_object* x_17 = _args[16]; lean_object* x_18 = _args[17]; lean_object* x_19 = _args[18]; lean_object* x_20 = _args[19]; +lean_object* x_21 = _args[20]; +lean_object* x_22 = _args[21]; _start: { -lean_object* x_21; -x_21 = l_Lean_Elab_Tactic_Omega_splitDisjunction___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, x_16, x_17, x_18, x_19, x_20); -return x_21; +uint8_t x_23; lean_object* x_24; +x_23 = lean_unbox(x_16); +lean_dec(x_16); +x_24 = l_Lean_Elab_Tactic_Omega_splitDisjunction___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, x_23, x_17, x_18, x_19, x_20, x_21, x_22); +return x_24; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___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, lean_object* x_16) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__6___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; _start: { -uint8_t x_17; lean_object* x_18; -x_17 = lean_unbox(x_8); +uint8_t x_19; lean_object* x_20; +x_19 = lean_unbox(x_12); +lean_dec(x_12); +x_20 = l_Lean_Elab_Tactic_Omega_splitDisjunction___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_19, x_13, x_14, x_15, x_16, x_17, x_18); +return x_20; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__7___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +_start: +{ +uint8_t x_19; uint8_t x_20; lean_object* x_21; +x_19 = lean_unbox(x_8); lean_dec(x_8); -x_18 = l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_17, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -return x_18; +x_20 = lean_unbox(x_12); +lean_dec(x_12); +x_21 = l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_19, x_9, x_10, x_11, x_20, x_13, x_14, x_15, x_16, x_17, x_18); +return x_21; } } -LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at_Lean_Elab_Tactic_Omega_omegaImpl___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_11; -x_11 = l_Lean_MVarId_assign___at_Lean_Elab_Tactic_Omega_omegaImpl___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +uint8_t x_13; lean_object* x_14; +x_13 = lean_unbox(x_6); +lean_dec(x_6); +x_14 = l_Lean_Elab_Tactic_Omega_splitDisjunction(x_1, x_2, x_3, x_4, x_5, x_13, x_7, x_8, x_9, x_10, x_11, x_12); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at_Lean_Elab_Tactic_Omega_omegaImpl___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; lean_object* x_14; +x_13 = lean_unbox(x_6); +lean_dec(x_6); +x_14 = l_Lean_MVarId_assign___at_Lean_Elab_Tactic_Omega_omegaImpl___spec__1(x_1, x_2, x_3, x_4, x_5, x_13, 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); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -return x_11; +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___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_7); +lean_dec(x_7); +x_15 = l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_14, x_8, x_9, x_10, x_11, x_12, x_13); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___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_8); +lean_dec(x_8); +x_16 = l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_15, x_9, x_10, x_11, x_12, x_13, x_14); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___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; lean_object* x_17; +x_16 = lean_unbox(x_9); +lean_dec(x_9); +x_17 = l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__3(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_13, x_14, x_15); +return x_17; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___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_8); +lean_dec(x_8); +x_16 = l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_15, x_9, x_10, x_11, x_12, x_13, x_14); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___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; lean_object* x_17; +x_16 = lean_unbox(x_9); +lean_dec(x_9); +x_17 = l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__5(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_13, x_14, x_15); +return x_17; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___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) { +_start: +{ +uint8_t x_13; lean_object* x_14; +x_13 = lean_unbox(x_6); +lean_dec(x_6); +x_14 = l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__6(x_1, x_2, x_3, x_4, x_5, x_13, x_7, x_8, x_9, x_10, x_11, x_12); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, 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_6); +lean_dec(x_6); +x_14 = l_Lean_Elab_Tactic_Omega_omegaImpl(x_1, x_2, x_3, x_4, x_5, x_13, x_7, x_8, x_9, x_10, x_11, x_12); +return x_14; } } LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_omega(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { @@ -30452,7 +31211,7 @@ lean_ctor_set(x_12, 0, x_10); lean_ctor_set(x_12, 1, x_1); lean_ctor_set(x_12, 2, x_9); lean_ctor_set(x_12, 3, x_11); -x_13 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_omegaImpl), 10, 2); +x_13 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_omegaImpl___boxed), 12, 2); lean_closure_set(x_13, 0, x_12); lean_closure_set(x_13, 1, x_2); x_14 = l_Lean_Elab_Tactic_Omega_OmegaM_run___rarg(x_13, x_3, x_4, x_5, x_6, x_7, x_8); @@ -32563,6 +33322,10 @@ l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1___closed lean_mark_persistent(l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1___closed__6); l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1___closed__7 = _init_l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1___closed__7(); lean_mark_persistent(l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1___closed__7); +l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1___closed__8 = _init_l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1___closed__8(); +lean_mark_persistent(l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1___closed__8); +l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1___closed__9 = _init_l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1___closed__9(); +lean_mark_persistent(l_panic___at_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___spec__1___closed__9); l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___lambda__2___closed__1 = _init_l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___lambda__2___closed__1(); lean_mark_persistent(l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___lambda__2___closed__1); l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___lambda__2___closed__2 = _init_l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___lambda__2___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/Tactic/Omega/OmegaM.c b/stage0/stdlib/Lean/Elab/Tactic/Omega/OmegaM.c index a1d5901307ad..c1ec87426a33 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Omega/OmegaM.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Omega/OmegaM.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Elab.Tactic.Omega.OmegaM -// Imports: Init.Omega.LinearCombo Init.Omega.Int Init.Omega.Logic Init.Data.BitVec.Basic Lean.Meta.AppBuilder +// Imports: Init.Omega.LinearCombo Init.Omega.Int Init.Omega.Logic Init.Data.BitVec.Basic Lean.Meta.AppBuilder Lean.Meta.Canonicalizer #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -14,16 +14,19 @@ extern "C" { #endif lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Elab_Tactic_Omega_lookup___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Omega_OmegaM_run___rarg___closed__2; +LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Elab_Tactic_Omega_lookup___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*); static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__83; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_OmegaM_run___rarg___lambda__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__22; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_lookup___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_intCast_x3f(lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__11; -LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___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_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__75; lean_object* l_Lean_mkNatLit(lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__27; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_atomsList___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_atomsList___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_Elab_Tactic_Omega_analyzeAtom___closed__79; LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_Elab_Tactic_Omega_lookup___spec__2(lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); @@ -31,13 +34,15 @@ LEAN_EXPORT uint8_t l_Lean_AssocList_contains___at_Lean_Elab_Tactic_Omega_lookup static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__81; static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__48; static lean_object* l_Lean_Elab_Tactic_Omega_groundNat_x3f___closed__15; -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(lean_object*, 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_Omega_lookup___spec__14(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__24; LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_Elab_Tactic_Omega_lookup___spec__2___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__3; +LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_OmegaM_run___spec__2(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__8; static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__73; static lean_object* l_Lean_Elab_Tactic_Omega_atomsCoeffs___rarg___closed__2; +LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_OmegaM_run___spec__2___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__99; static lean_object* l_Lean_Elab_Tactic_Omega_lookup___closed__3; lean_object* l_Lean_mkAppB(lean_object*, lean_object*, lean_object*); @@ -47,10 +52,11 @@ static lean_object* l_Lean_Elab_Tactic_Omega_groundInt_x3f___closed__4; static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__76; lean_object* l_Int_add___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_atomsCoeffs___rarg___closed__4; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___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_Omega_analyzeAtom___closed__4; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_lookup___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_Omega_analyzeAtom___closed__14; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_withoutModifyingState___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_Tactic_Omega_withoutModifyingState___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* l_Lean_MessageData_ofList(lean_object*); lean_object* l_Lean_PersistentArray_push___rarg(lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); @@ -69,18 +75,20 @@ static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__77; static lean_object* l_Lean_Elab_Tactic_Omega_groundNat_x3f___closed__14; static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__55; static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__80; +lean_object* l_Lean_Meta_Canonicalizer_CanonM_run___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkExpectedTypeHint(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_Elab_Tactic_Omega_lookup___spec__1(lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_atoms___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_atoms___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_AssocList_foldlM___at_Lean_Elab_Tactic_Omega_atoms___spec__2___boxed(lean_object*, lean_object*); lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_OmegaM_run___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__35; static lean_object* l_Lean_Elab_Tactic_Omega_atomsList___rarg___closed__1; lean_object* l_Int_ediv___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_groundInt_x3f_op(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__108; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_commitWhen___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_Omega_commitWhen___rarg___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*); static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__10; static lean_object* l_Lean_Elab_Tactic_Omega_groundNat_x3f___closed__8; static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__15; @@ -97,6 +105,7 @@ lean_object* l_Nat_sub___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_lookup___lambda__2___closed__4; lean_object* l_Int_sub___boxed(lean_object*, lean_object*); uint8_t lean_int_dec_le(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_OmegaM_run___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__26; static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__34; lean_object* l_Lean_Level_ofNat(lean_object*); @@ -135,24 +144,24 @@ static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__32; static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__78; lean_object* l_List_mapTR_loop___at_Lean_MessageData_instCoeListExprMessageData___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_groundNat_x3f___closed__6; -LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Elab_Tactic_Omega_lookup___spec__13(lean_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_Elab_Tactic_Omega_lookup___spec__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__98; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_atoms(lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_withoutModifyingState___rarg___closed__1; lean_object* l_Int_pow(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Elab_Tactic_Omega_atoms___spec__4(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14___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_Omega_lookup___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMap_insert___at_Lean_Elab_Tactic_Omega_lookup___spec__4(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__25; static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__46; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_cfg___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_cfg___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_Omega_groundNat_x3f(lean_object*); -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__101; lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__29; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_atoms___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_withoutModifyingState___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_Elab_Tactic_Omega_withoutModifyingState___rarg___lambda__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_st_mk_ref(lean_object*, lean_object*); lean_object* l_Lean_Meta_mkDecideProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Elab_Tactic_Omega_lookup___spec__8(lean_object*, lean_object*); @@ -163,7 +172,7 @@ static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__39; static lean_object* l_Lean_Elab_Tactic_Omega_groundInt_x3f___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_cfg___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_groundInt_x3f___closed__2; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_lookup___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_Omega_lookup___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__70; static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__12; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Omega_atoms___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -179,13 +188,15 @@ static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__7; static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__106; static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__23; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_lookup___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_Omega_lookup___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*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_OmegaM_run___rarg___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_mkEqReflWithExpectedType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__107; static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__21; lean_object* l_Nat_mul___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_atomsCoeffs___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_atomsList___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_OmegaM_run___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_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_atomsCoeffs___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_atomsList___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_natCast_x3f___closed__1; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Omega_lookup___spec__12(lean_object*, size_t, size_t, lean_object*); lean_object* l_Nat_add___boxed(lean_object*, lean_object*); @@ -200,16 +211,17 @@ lean_object* l_Lean_MessageData_ofExpr(lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__38; static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__60; static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__104; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_commitWhen___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_Elab_Tactic_Omega_commitWhen___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_EXPORT lean_object* l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___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*); static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__31; static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__90; static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__71; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_lookup(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_Omega_atomsCoeffs___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_lookup(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_Lean_Elab_Tactic_Omega_atomsCoeffs___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_Elab_Tactic_Omega_analyzeAtom___closed__58; static lean_object* l_Lean_Elab_Tactic_Omega_groundNat_x3f___closed__9; static lean_object* l_Lean_Elab_Tactic_Omega_groundNat_x3f___closed__2; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom(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_Omega_analyzeAtom(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__17; static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__65; @@ -229,6 +241,7 @@ static lean_object* l_Lean_Elab_Tactic_Omega_natCast_x3f___closed__2; LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Elab_Tactic_Omega_atoms___spec__4___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_cfg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_groundNat_x3f___closed__1; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_commitWhen___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* l_Lean_Meta_mkListLit(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__57; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); @@ -237,13 +250,13 @@ lean_object* l_Lean_Meta_mkEqRefl(lean_object*, lean_object*, lean_object*, lean static lean_object* l_Lean_Elab_Tactic_Omega_groundNat_x3f___closed__13; static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__54; LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Elab_Tactic_Omega_OmegaM_run___spec__1___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___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_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___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_Lean_Elab_Tactic_Omega_analyzeAtom___closed__40; static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__47; LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Elab_Tactic_Omega_atoms___spec__2(lean_object*, lean_object*); lean_object* l_Int_mul___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_groundInt_x3f(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_commitWhen___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_Tactic_Omega_commitWhen___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* 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_Tactic_Omega_analyzeAtom___closed__56; lean_object* l_Int_toNat(lean_object*); @@ -265,7 +278,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_atomsCoeffs(lean_object*); lean_object* l_Lean_Expr_int_x3f(lean_object*); lean_object* l_List_reverse___rarg(lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__89; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_lookup___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_Omega_lookup___lambda__1(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*); static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__9; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_atomsList___boxed(lean_object*); lean_object* l_Lean_mkHashMapImp___rarg(lean_object*); @@ -274,7 +287,7 @@ static lean_object* l_Lean_Elab_Tactic_Omega_groundNat_x3f___closed__11; size_t lean_usize_add(size_t, size_t); static lean_object* l_Lean_Elab_Tactic_Omega_groundNat_x3f___closed__4; lean_object* lean_array_uget(lean_object*, size_t); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_atoms___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_atoms___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_instToExprInt_mkNat(lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_lookup___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_atomsCoeffs___boxed(lean_object*); @@ -292,12 +305,14 @@ static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__36; static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__13; static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__61; lean_object* lean_array_get_size(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_cfg___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_cfg___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_int_neg(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_OmegaM_run___rarg___lambda__1___boxed(lean_object*, lean_object*, 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*); static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__5; 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_Elab_Tactic_Omega_withoutModifyingState___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_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Elab_Tactic_Omega_OmegaM_run___spec__1(lean_object*); lean_object* l_Lean_mkApp5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -310,15 +325,17 @@ static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__43; lean_object* l_Nat_pow___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashSet_toList___at_Lean_Elab_Tactic_Omega_lookup___spec__10(lean_object*); LEAN_EXPORT uint8_t l_Array_qsort_sort___at_Lean_Elab_Tactic_Omega_atoms___spec__4___lambda__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_OmegaM_run___spec__2___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_Elab_Tactic_Omega_groundInt_x3f___closed__3; LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Elab_Tactic_Omega_atoms___spec__4___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__64; lean_object* l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_lookup___closed__4; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_withoutModifyingState___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_withoutModifyingState___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_Lean_Elab_Tactic_Omega_analyzeAtom___closed__105; lean_object* l___private_Lean_Data_HashMap_0__Lean_numBucketsForCapacity(lean_object*); LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1(lean_object*, lean_object*); +lean_object* l_Lean_Meta_Canonicalizer_canon(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Elab_Tactic_Omega_State_atoms___default___spec__1(lean_object* x_1) { _start: { @@ -353,99 +370,229 @@ x_2 = l_Lean_mkHashMapImp___rarg(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic_Omega_OmegaM_run___rarg___closed__1() { +LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_OmegaM_run___spec__2___rarg(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) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(8u); -x_2 = l_Lean_mkHashMapImp___rarg(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_OmegaM_run___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_10; lean_object* x_11; +x_10 = lean_box(x_3); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_11 = lean_apply_7(x_1, x_10, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_11) == 0) { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_8 = lean_unsigned_to_nat(8u); -x_9 = l_Lean_mkHashMapImp___rarg(x_8); -x_10 = l_Lean_Elab_Tactic_Omega_OmegaM_run___rarg___closed__1; -x_11 = lean_st_mk_ref(x_10, x_7); +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* 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 = lean_st_mk_ref(x_9, 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_12); -lean_inc(x_15); -x_17 = lean_apply_8(x_1, x_15, x_12, x_2, x_3, x_4, x_5, x_6, x_16); -if (lean_obj_tag(x_17) == 0) +x_14 = lean_box(x_3); +x_15 = lean_apply_8(x_2, x_12, x_14, x_4, x_5, x_6, x_7, x_8, x_13); +return x_15; +} +else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_18 = lean_ctor_get(x_17, 0); +uint8_t 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_2); +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); -x_19 = lean_ctor_get(x_17, 1); -lean_inc(x_19); -lean_dec(x_17); -x_20 = lean_st_ref_get(x_15, x_19); -lean_dec(x_15); -x_21 = lean_ctor_get(x_20, 1); -lean_inc(x_21); -lean_dec(x_20); -x_22 = lean_st_ref_get(x_12, x_21); +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_ReaderT_bind___at_Lean_Elab_Tactic_Omega_OmegaM_run___spec__2(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_OmegaM_run___spec__2___rarg___boxed), 9, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_OmegaM_run___rarg___lambda__1(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_st_mk_ref(x_1, x_8); +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; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_OmegaM_run___rarg___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, 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; +x_12 = lean_st_mk_ref(x_1, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); lean_dec(x_12); -x_23 = !lean_is_exclusive(x_22); -if (x_23 == 0) +x_15 = lean_box(x_5); +lean_inc(x_4); +lean_inc(x_13); +x_16 = lean_apply_10(x_2, x_13, x_4, x_3, x_15, x_6, x_7, x_8, x_9, x_10, x_14); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_24; -x_24 = lean_ctor_get(x_22, 0); -lean_dec(x_24); -lean_ctor_set(x_22, 0, x_18); -return x_22; +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_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_st_ref_get(x_13, x_18); +lean_dec(x_13); +x_20 = lean_ctor_get(x_19, 1); +lean_inc(x_20); +lean_dec(x_19); +x_21 = lean_st_ref_get(x_4, x_20); +lean_dec(x_4); +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; +x_23 = lean_ctor_get(x_21, 0); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_17); +lean_ctor_set(x_24, 1, x_23); +lean_ctor_set(x_21, 0, x_24); +return x_21; } else { -lean_object* x_25; lean_object* x_26; -x_25 = lean_ctor_get(x_22, 1); +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_25 = lean_ctor_get(x_21, 0); +x_26 = lean_ctor_get(x_21, 1); +lean_inc(x_26); lean_inc(x_25); -lean_dec(x_22); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_18); -lean_ctor_set(x_26, 1, x_25); -return x_26; +lean_dec(x_21); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_17); +lean_ctor_set(x_27, 1, x_25); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_26); +return x_28; } } else { -uint8_t x_27; -lean_dec(x_15); -lean_dec(x_12); -x_27 = !lean_is_exclusive(x_17); -if (x_27 == 0) +uint8_t x_29; +lean_dec(x_13); +lean_dec(x_4); +x_29 = !lean_is_exclusive(x_16); +if (x_29 == 0) { -return x_17; +return x_16; } 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_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_16, 0); +x_31 = lean_ctor_get(x_16, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_16); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +return x_32; } } } } +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_OmegaM_run___rarg___lambda__3(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_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_8); +return x_10; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_Omega_OmegaM_run___rarg___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(8u); +x_2 = l_Lean_mkHashMapImp___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_Omega_OmegaM_run___rarg___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_OmegaM_run___rarg___lambda__3___boxed), 8, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_OmegaM_run___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; 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_8 = lean_unsigned_to_nat(8u); +x_9 = l_Lean_mkHashMapImp___rarg(x_8); +x_10 = l_Lean_Elab_Tactic_Omega_OmegaM_run___rarg___closed__1; +x_11 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_OmegaM_run___rarg___lambda__1___boxed), 8, 1); +lean_closure_set(x_11, 0, x_10); +x_12 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_OmegaM_run___rarg___lambda__2___boxed), 11, 3); +lean_closure_set(x_12, 0, x_9); +lean_closure_set(x_12, 1, x_1); +lean_closure_set(x_12, 2, x_2); +x_13 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_OmegaM_run___spec__2___rarg___boxed), 9, 2); +lean_closure_set(x_13, 0, x_11); +lean_closure_set(x_13, 1, x_12); +x_14 = l_Lean_Elab_Tactic_Omega_OmegaM_run___rarg___closed__2; +x_15 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_OmegaM_run___spec__2___rarg___boxed), 9, 2); +lean_closure_set(x_15, 0, x_13); +lean_closure_set(x_15, 1, x_14); +x_16 = 3; +x_17 = l_Lean_Meta_Canonicalizer_CanonM_run___rarg(x_15, x_16, x_3, x_4, x_5, x_6, x_7); +return x_17; +} +} LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_OmegaM_run(lean_object* x_1) { _start: { @@ -463,34 +610,88 @@ lean_dec(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_cfg___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_ReaderT_bind___at_Lean_Elab_Tactic_Omega_OmegaM_run___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) { _start: { -lean_object* x_7; -x_7 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_7, 0, x_1); -lean_ctor_set(x_7, 1, x_6); -return x_7; +uint8_t x_10; lean_object* x_11; +x_10 = lean_unbox(x_3); +lean_dec(x_3); +x_11 = l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_OmegaM_run___spec__2___rarg(x_1, x_2, x_10, x_4, x_5, x_6, x_7, x_8, x_9); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_OmegaM_run___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, 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_Elab_Tactic_Omega_OmegaM_run___rarg___lambda__1(x_1, x_9, 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_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_OmegaM_run___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) { +_start: +{ +uint8_t x_12; lean_object* x_13; +x_12 = lean_unbox(x_5); +lean_dec(x_5); +x_13 = l_Lean_Elab_Tactic_Omega_OmegaM_run___rarg___lambda__2(x_1, x_2, x_3, x_4, x_12, x_6, x_7, x_8, x_9, x_10, x_11); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_OmegaM_run___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) { +_start: +{ +uint8_t x_9; lean_object* x_10; +x_9 = lean_unbox(x_2); +lean_dec(x_2); +x_10 = l_Lean_Elab_Tactic_Omega_OmegaM_run___rarg___lambda__3(x_1, x_9, 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_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_cfg___rarg(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_9; +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_Elab_Tactic_Omega_cfg(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_cfg___rarg___boxed), 6, 0); +x_3 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_cfg___rarg___boxed), 8, 0); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_cfg___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_Elab_Tactic_Omega_cfg___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_7; -x_7 = l_Lean_Elab_Tactic_Omega_cfg___rarg(x_1, x_2, x_3, x_4, x_5, x_6); +uint8_t x_9; lean_object* x_10; +x_9 = lean_unbox(x_2); +lean_dec(x_2); +x_10 = l_Lean_Elab_Tactic_Omega_cfg___rarg(x_1, x_9, 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_7; +return x_10; } } LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_cfg___boxed(lean_object* x_1, lean_object* x_2) { @@ -691,57 +892,57 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_atoms___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_Elab_Tactic_Omega_atoms___rarg(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) { _start: { -lean_object* x_8; uint8_t x_9; -x_8 = lean_st_ref_get(x_1, x_7); -x_9 = !lean_is_exclusive(x_8); -if (x_9 == 0) +lean_object* x_10; uint8_t x_11; +x_10 = lean_st_ref_get(x_1, x_9); +x_11 = !lean_is_exclusive(x_10); +if (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; size_t x_18; size_t x_19; lean_object* x_20; -x_10 = lean_ctor_get(x_8, 0); -x_11 = l_Lean_HashMap_toArray___at_Lean_Elab_Tactic_Omega_atoms___spec__1(x_10); -x_12 = lean_array_get_size(x_11); -x_13 = lean_unsigned_to_nat(1u); -x_14 = lean_nat_sub(x_12, x_13); -lean_dec(x_12); -x_15 = lean_unsigned_to_nat(0u); -x_16 = l_Array_qsort_sort___at_Lean_Elab_Tactic_Omega_atoms___spec__4(x_11, x_15, x_14); +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; size_t x_20; size_t x_21; lean_object* x_22; +x_12 = lean_ctor_get(x_10, 0); +x_13 = l_Lean_HashMap_toArray___at_Lean_Elab_Tactic_Omega_atoms___spec__1(x_12); +x_14 = lean_array_get_size(x_13); +x_15 = lean_unsigned_to_nat(1u); +x_16 = lean_nat_sub(x_14, x_15); lean_dec(x_14); -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_Tactic_Omega_atoms___spec__5(x_18, x_19, x_16); -lean_ctor_set(x_8, 0, x_20); -return x_8; +x_17 = lean_unsigned_to_nat(0u); +x_18 = l_Array_qsort_sort___at_Lean_Elab_Tactic_Omega_atoms___spec__4(x_13, x_17, x_16); +lean_dec(x_16); +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_Lean_Elab_Tactic_Omega_atoms___spec__5(x_20, x_21, x_18); +lean_ctor_set(x_10, 0, x_22); +return x_10; } 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; size_t x_30; size_t x_31; lean_object* x_32; lean_object* x_33; -x_21 = lean_ctor_get(x_8, 0); -x_22 = lean_ctor_get(x_8, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_8); -x_23 = l_Lean_HashMap_toArray___at_Lean_Elab_Tactic_Omega_atoms___spec__1(x_21); -x_24 = lean_array_get_size(x_23); -x_25 = lean_unsigned_to_nat(1u); -x_26 = lean_nat_sub(x_24, x_25); -lean_dec(x_24); -x_27 = lean_unsigned_to_nat(0u); -x_28 = l_Array_qsort_sort___at_Lean_Elab_Tactic_Omega_atoms___spec__4(x_23, x_27, x_26); +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; size_t x_32; size_t x_33; lean_object* x_34; lean_object* x_35; +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 = l_Lean_HashMap_toArray___at_Lean_Elab_Tactic_Omega_atoms___spec__1(x_23); +x_26 = lean_array_get_size(x_25); +x_27 = lean_unsigned_to_nat(1u); +x_28 = lean_nat_sub(x_26, x_27); lean_dec(x_26); -x_29 = lean_array_get_size(x_28); -x_30 = lean_usize_of_nat(x_29); -lean_dec(x_29); -x_31 = 0; -x_32 = l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_Omega_atoms___spec__5(x_30, x_31, x_28); -x_33 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_33, 1, x_22); -return x_33; +x_29 = lean_unsigned_to_nat(0u); +x_30 = l_Array_qsort_sort___at_Lean_Elab_Tactic_Omega_atoms___spec__4(x_25, x_29, x_28); +lean_dec(x_28); +x_31 = lean_array_get_size(x_30); +x_32 = lean_usize_of_nat(x_31); +lean_dec(x_31); +x_33 = 0; +x_34 = l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_Omega_atoms___spec__5(x_32, x_33, x_30); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_24); +return x_35; } } } @@ -749,7 +950,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_atoms(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_atoms___rarg___boxed), 7, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_atoms___rarg___boxed), 9, 0); return x_2; } } @@ -807,18 +1008,21 @@ x_6 = l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_Omega_atoms___spec__5(x_4, x_ return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_atoms___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_Tactic_Omega_atoms___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_8; -x_8 = l_Lean_Elab_Tactic_Omega_atoms___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +uint8_t x_10; lean_object* x_11; +x_10 = lean_unbox(x_3); +lean_dec(x_3); +x_11 = l_Lean_Elab_Tactic_Omega_atoms___rarg(x_1, x_2, x_10, 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); lean_dec(x_1); -return x_8; +return x_11; } } LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_atoms___boxed(lean_object* x_1) { @@ -858,38 +1062,41 @@ x_3 = l_Lean_Expr_const___override(x_2, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_atomsList___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_Elab_Tactic_Omega_atomsList___rarg(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) { _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; -x_8 = l_Lean_Elab_Tactic_Omega_atoms___rarg(x_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); -lean_inc(x_10); -lean_dec(x_8); -x_11 = lean_array_to_list(lean_box(0), x_9); -x_12 = l_Lean_Elab_Tactic_Omega_atomsList___rarg___closed__3; -x_13 = l_Lean_Meta_mkListLit(x_12, x_11, x_3, x_4, x_5, x_6, x_10); -return x_13; +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_10 = l_Lean_Elab_Tactic_Omega_atoms___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +x_13 = lean_array_to_list(lean_box(0), x_11); +x_14 = l_Lean_Elab_Tactic_Omega_atomsList___rarg___closed__3; +x_15 = l_Lean_Meta_mkListLit(x_14, x_13, x_5, x_6, x_7, x_8, x_12); +return x_15; } } LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_atomsList(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_atomsList___rarg___boxed), 7, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_atomsList___rarg___boxed), 9, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_atomsList___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_Tactic_Omega_atomsList___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_8; -x_8 = l_Lean_Elab_Tactic_Omega_atomsList___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +uint8_t x_10; lean_object* x_11; +x_10 = lean_unbox(x_3); +lean_dec(x_3); +x_11 = l_Lean_Elab_Tactic_Omega_atomsList___rarg(x_1, x_2, x_10, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -return x_8; +return x_11; } } LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_atomsList___boxed(lean_object* x_1) { @@ -955,60 +1162,60 @@ x_3 = l_Lean_Expr_const___override(x_2, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_atomsCoeffs___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_Elab_Tactic_Omega_atomsCoeffs___rarg(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) { _start: { -lean_object* x_8; -x_8 = l_Lean_Elab_Tactic_Omega_atomsList___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_8) == 0) +lean_object* x_10; +x_10 = l_Lean_Elab_Tactic_Omega_atomsList___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_10) == 0) { -uint8_t x_9; -x_9 = !lean_is_exclusive(x_8); -if (x_9 == 0) +uint8_t x_11; +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = lean_ctor_get(x_8, 0); -x_11 = l_Lean_Elab_Tactic_Omega_atomsCoeffs___rarg___closed__6; -x_12 = l_Lean_Expr_app___override(x_11, x_10); -lean_ctor_set(x_8, 0, x_12); -return x_8; +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_10, 0); +x_13 = l_Lean_Elab_Tactic_Omega_atomsCoeffs___rarg___closed__6; +x_14 = l_Lean_Expr_app___override(x_13, x_12); +lean_ctor_set(x_10, 0, x_14); +return x_10; } 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_8, 0); -x_14 = lean_ctor_get(x_8, 1); -lean_inc(x_14); -lean_inc(x_13); -lean_dec(x_8); -x_15 = l_Lean_Elab_Tactic_Omega_atomsCoeffs___rarg___closed__6; -x_16 = l_Lean_Expr_app___override(x_15, x_13); -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; +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_10, 0); +x_16 = lean_ctor_get(x_10, 1); +lean_inc(x_16); +lean_inc(x_15); +lean_dec(x_10); +x_17 = l_Lean_Elab_Tactic_Omega_atomsCoeffs___rarg___closed__6; +x_18 = l_Lean_Expr_app___override(x_17, x_15); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_16); +return x_19; } } else { -uint8_t x_18; -x_18 = !lean_is_exclusive(x_8); -if (x_18 == 0) +uint8_t x_20; +x_20 = !lean_is_exclusive(x_10); +if (x_20 == 0) { -return x_8; +return x_10; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_8, 0); -x_20 = lean_ctor_get(x_8, 1); -lean_inc(x_20); -lean_inc(x_19); -lean_dec(x_8); -x_21 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_21, 0, x_19); -lean_ctor_set(x_21, 1, x_20); -return x_21; +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_10, 0); +x_22 = lean_ctor_get(x_10, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_10); +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; } } } @@ -1017,18 +1224,21 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_atomsCoeffs(lean_object* x_1) _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_atomsCoeffs___rarg___boxed), 7, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_atomsCoeffs___rarg___boxed), 9, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_atomsCoeffs___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_Tactic_Omega_atomsCoeffs___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_8; -x_8 = l_Lean_Elab_Tactic_Omega_atomsCoeffs___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +uint8_t x_10; lean_object* x_11; +x_10 = lean_unbox(x_3); +lean_dec(x_3); +x_11 = l_Lean_Elab_Tactic_Omega_atomsCoeffs___rarg(x_1, x_2, x_10, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -return x_8; +return x_11; } } LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_atomsCoeffs___boxed(lean_object* x_1) { @@ -1040,148 +1250,149 @@ lean_dec(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_commitWhen___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_commitWhen___rarg___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_object* x_11, lean_object* x_12) { _start: { -lean_object* x_11; -x_11 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_11, 0, x_1); -lean_ctor_set(x_11, 1, x_10); -return x_11; +lean_object* x_13; +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_1); +lean_ctor_set(x_13, 1, x_12); +return x_13; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_commitWhen___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_Elab_Tactic_Omega_commitWhen___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_10 = lean_st_ref_get(x_3, x_9); -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_10, 1); -lean_inc(x_12); -lean_dec(x_10); -x_13 = lean_st_ref_get(x_2, x_12); -x_14 = lean_ctor_get(x_13, 0); +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_12 = lean_st_ref_get(x_3, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); +lean_dec(x_12); +x_15 = lean_st_ref_get(x_2, 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_box(x_5); lean_inc(x_3); lean_inc(x_2); -x_16 = lean_apply_8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_15); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_17, 1); -lean_inc(x_18); -x_19 = lean_unbox(x_18); -lean_dec(x_18); -if (x_19 == 0) +x_19 = lean_apply_10(x_1, x_2, x_3, x_4, x_18, x_6, x_7, x_8, x_9, x_10, 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; 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_20 = lean_ctor_get(x_16, 1); +lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_20 = lean_ctor_get(x_19, 0); lean_inc(x_20); -lean_dec(x_16); -x_21 = lean_ctor_get(x_17, 0); +x_21 = lean_ctor_get(x_20, 1); lean_inc(x_21); -lean_dec(x_17); -x_22 = lean_st_ref_take(x_3, x_20); -x_23 = lean_ctor_get(x_22, 1); +x_22 = lean_unbox(x_21); +lean_dec(x_21); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_23 = lean_ctor_get(x_19, 1); lean_inc(x_23); -lean_dec(x_22); -x_24 = lean_st_ref_set(x_3, x_11, x_23); +lean_dec(x_19); +x_24 = lean_ctor_get(x_20, 0); +lean_inc(x_24); +lean_dec(x_20); +x_25 = lean_st_ref_take(x_3, x_23); +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +lean_dec(x_25); +x_27 = lean_st_ref_set(x_3, x_13, x_26); lean_dec(x_3); -x_25 = lean_ctor_get(x_24, 1); -lean_inc(x_25); -lean_dec(x_24); -x_26 = lean_st_ref_take(x_2, x_25); -x_27 = lean_ctor_get(x_26, 1); -lean_inc(x_27); -lean_dec(x_26); -x_28 = lean_st_ref_set(x_2, x_14, x_27); +x_28 = lean_ctor_get(x_27, 1); +lean_inc(x_28); +lean_dec(x_27); +x_29 = lean_st_ref_take(x_2, x_28); +x_30 = lean_ctor_get(x_29, 1); +lean_inc(x_30); +lean_dec(x_29); +x_31 = lean_st_ref_set(x_2, x_16, x_30); lean_dec(x_2); -x_29 = !lean_is_exclusive(x_28); -if (x_29 == 0) +x_32 = !lean_is_exclusive(x_31); +if (x_32 == 0) { -lean_object* x_30; -x_30 = lean_ctor_get(x_28, 0); -lean_dec(x_30); -lean_ctor_set(x_28, 0, x_21); -return x_28; +lean_object* x_33; +x_33 = lean_ctor_get(x_31, 0); +lean_dec(x_33); +lean_ctor_set(x_31, 0, x_24); +return x_31; } else { -lean_object* x_31; lean_object* x_32; -x_31 = lean_ctor_get(x_28, 1); -lean_inc(x_31); -lean_dec(x_28); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_21); -lean_ctor_set(x_32, 1, x_31); -return x_32; +lean_object* x_34; lean_object* x_35; +x_34 = lean_ctor_get(x_31, 1); +lean_inc(x_34); +lean_dec(x_31); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_24); +lean_ctor_set(x_35, 1, x_34); +return x_35; } } else { -uint8_t x_33; -lean_dec(x_14); -lean_dec(x_11); +uint8_t x_36; +lean_dec(x_16); +lean_dec(x_13); lean_dec(x_3); lean_dec(x_2); -x_33 = !lean_is_exclusive(x_16); -if (x_33 == 0) +x_36 = !lean_is_exclusive(x_19); +if (x_36 == 0) { -lean_object* x_34; lean_object* x_35; -x_34 = lean_ctor_get(x_16, 0); -lean_dec(x_34); -x_35 = lean_ctor_get(x_17, 0); -lean_inc(x_35); -lean_dec(x_17); -lean_ctor_set(x_16, 0, x_35); -return x_16; +lean_object* x_37; lean_object* x_38; +x_37 = lean_ctor_get(x_19, 0); +lean_dec(x_37); +x_38 = lean_ctor_get(x_20, 0); +lean_inc(x_38); +lean_dec(x_20); +lean_ctor_set(x_19, 0, x_38); +return x_19; } else { -lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_36 = lean_ctor_get(x_16, 1); -lean_inc(x_36); -lean_dec(x_16); -x_37 = lean_ctor_get(x_17, 0); -lean_inc(x_37); -lean_dec(x_17); -x_38 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_38, 0, x_37); -lean_ctor_set(x_38, 1, x_36); -return x_38; +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_19, 1); +lean_inc(x_39); +lean_dec(x_19); +x_40 = lean_ctor_get(x_20, 0); +lean_inc(x_40); +lean_dec(x_20); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_39); +return x_41; } } } else { -uint8_t x_39; -lean_dec(x_14); -lean_dec(x_11); +uint8_t x_42; +lean_dec(x_16); +lean_dec(x_13); lean_dec(x_3); lean_dec(x_2); -x_39 = !lean_is_exclusive(x_16); -if (x_39 == 0) +x_42 = !lean_is_exclusive(x_19); +if (x_42 == 0) { -return x_16; +return x_19; } else { -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_16, 0); -x_41 = lean_ctor_get(x_16, 1); -lean_inc(x_41); -lean_inc(x_40); -lean_dec(x_16); -x_42 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_42, 0, x_40); -lean_ctor_set(x_42, 1, x_41); -return x_42; +lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_43 = lean_ctor_get(x_19, 0); +x_44 = lean_ctor_get(x_19, 1); +lean_inc(x_44); +lean_inc(x_43); +lean_dec(x_19); +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; } } } @@ -1190,77 +1401,94 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_commitWhen(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_commitWhen___rarg), 9, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_commitWhen___rarg___boxed), 11, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_commitWhen___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_commitWhen___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_11; -x_11 = l_Lean_Elab_Tactic_Omega_commitWhen___rarg___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +uint8_t x_13; lean_object* x_14; +x_13 = lean_unbox(x_6); +lean_dec(x_6); +x_14 = l_Lean_Elab_Tactic_Omega_commitWhen___rarg___lambda__1(x_1, x_2, x_3, x_4, x_5, x_13, 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); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_11; +return x_14; } } -LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___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_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_commitWhen___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) { _start: { -lean_object* x_11; +uint8_t x_12; lean_object* x_13; +x_12 = lean_unbox(x_5); +lean_dec(x_5); +x_13 = l_Lean_Elab_Tactic_Omega_commitWhen___rarg(x_1, x_2, x_3, x_4, x_12, x_6, x_7, x_8, x_9, x_10, x_11); +return x_13; +} +} +LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg(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_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_box(x_6); +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); lean_inc(x_3); -x_11 = lean_apply_8(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_11) == 0) +x_14 = lean_apply_10(x_1, x_3, x_4, x_5, x_13, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_14) == 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 = lean_apply_9(x_2, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); -return x_14; +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* 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 = lean_box(x_6); +x_18 = lean_apply_11(x_2, x_15, x_3, x_4, x_5, x_17, x_7, x_8, x_9, x_10, x_11, x_16); +return x_18; } else { -uint8_t x_15; +uint8_t x_19; +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_15 = !lean_is_exclusive(x_11); -if (x_15 == 0) +x_19 = !lean_is_exclusive(x_14); +if (x_19 == 0) { -return x_11; +return x_14; } else { -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; +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; } } } @@ -1269,66 +1497,89 @@ LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModif _start: { lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg), 10, 0); +x_3 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg___boxed), 12, 0); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_withoutModifyingState___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_withoutModifyingState___rarg___lambda__1(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) { _start: { -uint8_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_10 = 0; -x_11 = lean_box(x_10); -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(0, 2, 0); -lean_ctor_set(x_13, 0, x_12); -lean_ctor_set(x_13, 1, x_9); -return x_13; +uint8_t x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_12 = 0; +x_13 = lean_box(x_12); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_1); +lean_ctor_set(x_14, 1, x_13); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_11); +return x_15; } } static lean_object* _init_l_Lean_Elab_Tactic_Omega_withoutModifyingState___rarg___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_withoutModifyingState___rarg___lambda__1___boxed), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_withoutModifyingState___rarg___lambda__1___boxed), 11, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_withoutModifyingState___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_Elab_Tactic_Omega_withoutModifyingState___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = l_Lean_Elab_Tactic_Omega_withoutModifyingState___rarg___closed__1; -x_11 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg), 10, 2); -lean_closure_set(x_11, 0, x_1); -lean_closure_set(x_11, 1, x_10); -x_12 = l_Lean_Elab_Tactic_Omega_commitWhen___rarg(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_12; +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = l_Lean_Elab_Tactic_Omega_withoutModifyingState___rarg___closed__1; +x_13 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg___boxed), 12, 2); +lean_closure_set(x_13, 0, x_1); +lean_closure_set(x_13, 1, x_12); +x_14 = l_Lean_Elab_Tactic_Omega_commitWhen___rarg(x_13, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_14; } } LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_withoutModifyingState(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_withoutModifyingState___rarg), 9, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_withoutModifyingState___rarg___boxed), 11, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_withoutModifyingState___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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; -x_10 = l_Lean_Elab_Tactic_Omega_withoutModifyingState___rarg___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +uint8_t x_13; lean_object* x_14; +x_13 = lean_unbox(x_6); +lean_dec(x_6); +x_14 = l_ReaderT_bind___at_Lean_Elab_Tactic_Omega_withoutModifyingState___spec__1___rarg(x_1, x_2, x_3, x_4, x_5, x_13, x_7, x_8, x_9, x_10, x_11, x_12); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_withoutModifyingState___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, 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_5); +lean_dec(x_5); +x_13 = l_Lean_Elab_Tactic_Omega_withoutModifyingState___rarg___lambda__1(x_1, x_2, x_3, x_4, x_12, 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_10; +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_withoutModifyingState___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) { +_start: +{ +uint8_t x_12; lean_object* x_13; +x_12 = lean_unbox(x_5); +lean_dec(x_5); +x_13 = l_Lean_Elab_Tactic_Omega_withoutModifyingState___rarg(x_1, x_2, x_3, x_4, x_12, x_6, x_7, x_8, x_9, x_10, x_11); +return x_13; } } static lean_object* _init_l_Lean_Elab_Tactic_Omega_natCast_x3f___closed__1() { @@ -3773,1114 +4024,1092 @@ x_3 = l_Lean_Expr_const___override(x_2, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom(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) { _start: { -lean_object* x_10; lean_object* x_11; -x_10 = l_Lean_Expr_getAppFnArgs(x_1); -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -if (lean_obj_tag(x_11) == 1) +lean_object* x_12; lean_object* x_13; +x_12 = l_Lean_Expr_getAppFnArgs(x_1); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +if (lean_obj_tag(x_13) == 1) { -lean_object* x_12; -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -switch (lean_obj_tag(x_12)) { +lean_object* x_14; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +switch (lean_obj_tag(x_14)) { case 0: { -lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t 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); -x_13 = lean_ctor_get(x_10, 1); -lean_inc(x_13); -lean_dec(x_10); -x_14 = lean_ctor_get(x_11, 1); -lean_inc(x_14); -lean_dec(x_11); -x_15 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__2; -x_16 = lean_string_dec_eq(x_14, x_15); -lean_dec(x_14); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; +x_15 = lean_ctor_get(x_12, 1); +lean_inc(x_15); +lean_dec(x_12); +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); lean_dec(x_13); -x_17 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_9); -return x_18; +x_17 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__2; +x_18 = lean_string_dec_eq(x_16, x_17); +lean_dec(x_16); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +lean_dec(x_15); +x_19 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_11); +return x_20; } else { -lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_19 = lean_array_get_size(x_13); -x_20 = lean_unsigned_to_nat(5u); -x_21 = lean_nat_dec_eq(x_19, x_20); -lean_dec(x_19); -if (x_21 == 0) +lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_21 = lean_array_get_size(x_15); +x_22 = lean_unsigned_to_nat(5u); +x_23 = lean_nat_dec_eq(x_21, x_22); +lean_dec(x_21); +if (x_23 == 0) { -lean_object* x_22; lean_object* x_23; -lean_dec(x_13); -x_22 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_9); -return x_23; +lean_object* x_24; lean_object* x_25; +lean_dec(x_15); +x_24 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_11); +return x_25; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; -x_24 = lean_unsigned_to_nat(0u); -x_25 = lean_array_fget(x_13, x_24); -x_26 = lean_unsigned_to_nat(1u); -x_27 = lean_array_fget(x_13, x_26); -x_28 = lean_unsigned_to_nat(2u); -x_29 = lean_array_fget(x_13, x_28); -x_30 = lean_unsigned_to_nat(3u); -x_31 = lean_array_fget(x_13, x_30); -x_32 = lean_unsigned_to_nat(4u); -x_33 = lean_array_fget(x_13, x_32); -lean_dec(x_13); -x_34 = l_Lean_Elab_Tactic_Omega_atomsList___rarg___closed__3; -x_35 = lean_expr_eqv(x_25, x_34); -if (x_35 == 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; uint8_t x_37; +x_26 = lean_unsigned_to_nat(0u); +x_27 = lean_array_fget(x_15, x_26); +x_28 = lean_unsigned_to_nat(1u); +x_29 = lean_array_fget(x_15, x_28); +x_30 = lean_unsigned_to_nat(2u); +x_31 = lean_array_fget(x_15, x_30); +x_32 = lean_unsigned_to_nat(3u); +x_33 = lean_array_fget(x_15, x_32); +x_34 = lean_unsigned_to_nat(4u); +x_35 = lean_array_fget(x_15, x_34); +lean_dec(x_15); +x_36 = l_Lean_Elab_Tactic_Omega_atomsList___rarg___closed__3; +x_37 = lean_expr_eqv(x_27, x_36); +if (x_37 == 0) { -lean_object* x_36; lean_object* x_37; +lean_object* x_38; lean_object* x_39; +lean_dec(x_35); lean_dec(x_33); lean_dec(x_31); lean_dec(x_29); lean_dec(x_27); -lean_dec(x_25); -x_36 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_36); -lean_ctor_set(x_37, 1, x_9); -return x_37; -} -else -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_38 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__7; -x_39 = l_Lean_mkApp5(x_38, x_25, x_27, x_29, x_31, x_33); -x_40 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_41 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_40, x_39); -x_42 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_9); -return x_42; +x_38 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_11); +return x_39; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_40 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__7; +x_41 = l_Lean_mkApp5(x_40, x_27, x_29, x_31, x_33, x_35); +x_42 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_43 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_42, x_41); +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_44, 1, x_11); +return x_44; } } } } case 1: { -lean_object* x_43; -x_43 = lean_ctor_get(x_12, 0); -lean_inc(x_43); -if (lean_obj_tag(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_10, 1); -lean_inc(x_44); -lean_dec(x_10); -x_45 = lean_ctor_get(x_11, 1); +lean_object* x_45; +x_45 = lean_ctor_get(x_14, 0); lean_inc(x_45); -lean_dec(x_11); +if (lean_obj_tag(x_45) == 0) +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; uint8_t x_50; x_46 = lean_ctor_get(x_12, 1); lean_inc(x_46); lean_dec(x_12); -x_47 = l_Lean_Elab_Tactic_Omega_natCast_x3f___closed__1; -x_48 = lean_string_dec_eq(x_46, x_47); -if (x_48 == 0) -{ -lean_object* x_49; uint8_t x_50; -x_49 = l_Lean_Elab_Tactic_Omega_groundNat_x3f___closed__4; -x_50 = lean_string_dec_eq(x_46, x_49); +x_47 = lean_ctor_get(x_13, 1); +lean_inc(x_47); +lean_dec(x_13); +x_48 = lean_ctor_get(x_14, 1); +lean_inc(x_48); +lean_dec(x_14); +x_49 = l_Lean_Elab_Tactic_Omega_natCast_x3f___closed__1; +x_50 = lean_string_dec_eq(x_48, x_49); if (x_50 == 0) { lean_object* x_51; uint8_t x_52; -x_51 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__8; -x_52 = lean_string_dec_eq(x_46, x_51); +x_51 = l_Lean_Elab_Tactic_Omega_groundNat_x3f___closed__4; +x_52 = lean_string_dec_eq(x_48, x_51); if (x_52 == 0) { lean_object* x_53; uint8_t x_54; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -x_53 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__9; -x_54 = lean_string_dec_eq(x_46, x_53); +x_53 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__8; +x_54 = lean_string_dec_eq(x_48, x_53); if (x_54 == 0) { lean_object* x_55; uint8_t x_56; -x_55 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__10; -x_56 = lean_string_dec_eq(x_46, x_55); -lean_dec(x_46); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +x_55 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__9; +x_56 = lean_string_dec_eq(x_48, x_55); if (x_56 == 0) { -lean_object* x_57; lean_object* x_58; -lean_dec(x_45); -lean_dec(x_44); -x_57 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_58 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_58, 0, x_57); -lean_ctor_set(x_58, 1, x_9); -return x_58; +lean_object* x_57; uint8_t x_58; +x_57 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__10; +x_58 = lean_string_dec_eq(x_48, x_57); +lean_dec(x_48); +if (x_58 == 0) +{ +lean_object* x_59; lean_object* x_60; +lean_dec(x_47); +lean_dec(x_46); +x_59 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_60 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_60, 0, x_59); +lean_ctor_set(x_60, 1, x_11); +return x_60; } else { -lean_object* x_59; uint8_t x_60; -x_59 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__11; -x_60 = lean_string_dec_eq(x_45, x_59); -lean_dec(x_45); -if (x_60 == 0) +lean_object* x_61; uint8_t x_62; +x_61 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__11; +x_62 = lean_string_dec_eq(x_47, x_61); +lean_dec(x_47); +if (x_62 == 0) { -lean_object* x_61; lean_object* x_62; -lean_dec(x_44); -x_61 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_62 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_62, 0, x_61); -lean_ctor_set(x_62, 1, x_9); -return x_62; +lean_object* x_63; lean_object* x_64; +lean_dec(x_46); +x_63 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_64 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_64, 0, x_63); +lean_ctor_set(x_64, 1, x_11); +return x_64; } else { -lean_object* x_63; lean_object* x_64; uint8_t x_65; -x_63 = lean_array_get_size(x_44); -x_64 = lean_unsigned_to_nat(4u); -x_65 = lean_nat_dec_eq(x_63, x_64); -lean_dec(x_63); -if (x_65 == 0) +lean_object* x_65; lean_object* x_66; uint8_t x_67; +x_65 = lean_array_get_size(x_46); +x_66 = lean_unsigned_to_nat(4u); +x_67 = lean_nat_dec_eq(x_65, x_66); +lean_dec(x_65); +if (x_67 == 0) { -lean_object* x_66; lean_object* x_67; -lean_dec(x_44); -x_66 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_67 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_67, 0, x_66); -lean_ctor_set(x_67, 1, x_9); -return x_67; +lean_object* x_68; lean_object* x_69; +lean_dec(x_46); +x_68 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_69 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_69, 0, x_68); +lean_ctor_set(x_69, 1, x_11); +return x_69; } else { -lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; -x_68 = lean_unsigned_to_nat(2u); -x_69 = lean_array_fget(x_44, x_68); -x_70 = lean_unsigned_to_nat(3u); -x_71 = lean_array_fget(x_44, x_70); -lean_dec(x_44); -x_72 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__14; +lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_70 = lean_unsigned_to_nat(2u); +x_71 = lean_array_fget(x_46, x_70); +x_72 = lean_unsigned_to_nat(3u); +x_73 = lean_array_fget(x_46, x_72); +lean_dec(x_46); +x_74 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__14; +lean_inc(x_73); lean_inc(x_71); -lean_inc(x_69); -x_73 = l_Lean_mkAppB(x_72, x_69, x_71); -x_74 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__17; -x_75 = l_Lean_mkAppB(x_74, x_69, x_71); -x_76 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_77 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_76, x_73); -x_78 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_77, x_75); -x_79 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_79, 0, x_78); -lean_ctor_set(x_79, 1, x_9); -return x_79; +x_75 = l_Lean_mkAppB(x_74, x_71, x_73); +x_76 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__17; +x_77 = l_Lean_mkAppB(x_76, x_71, x_73); +x_78 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_79 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_78, x_75); +x_80 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_79, x_77); +x_81 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_81, 0, x_80); +lean_ctor_set(x_81, 1, x_11); +return x_81; } } } } else { -lean_object* x_80; uint8_t x_81; -lean_dec(x_46); -x_80 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__18; -x_81 = lean_string_dec_eq(x_45, x_80); -lean_dec(x_45); -if (x_81 == 0) +lean_object* x_82; uint8_t x_83; +lean_dec(x_48); +x_82 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__18; +x_83 = lean_string_dec_eq(x_47, x_82); +lean_dec(x_47); +if (x_83 == 0) { -lean_object* x_82; lean_object* x_83; -lean_dec(x_44); -x_82 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_83 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_83, 0, x_82); -lean_ctor_set(x_83, 1, x_9); -return x_83; +lean_object* x_84; lean_object* x_85; +lean_dec(x_46); +x_84 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_85 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_85, 0, x_84); +lean_ctor_set(x_85, 1, x_11); +return x_85; } else { -lean_object* x_84; lean_object* x_85; uint8_t x_86; -x_84 = lean_array_get_size(x_44); -x_85 = lean_unsigned_to_nat(4u); -x_86 = lean_nat_dec_eq(x_84, x_85); -lean_dec(x_84); -if (x_86 == 0) +lean_object* x_86; lean_object* x_87; uint8_t x_88; +x_86 = lean_array_get_size(x_46); +x_87 = lean_unsigned_to_nat(4u); +x_88 = lean_nat_dec_eq(x_86, x_87); +lean_dec(x_86); +if (x_88 == 0) { -lean_object* x_87; lean_object* x_88; -lean_dec(x_44); -x_87 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_88 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_88, 0, x_87); -lean_ctor_set(x_88, 1, x_9); -return x_88; +lean_object* x_89; lean_object* x_90; +lean_dec(x_46); +x_89 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_90 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_90, 0, x_89); +lean_ctor_set(x_90, 1, x_11); +return x_90; } else { -lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; -x_89 = lean_unsigned_to_nat(2u); -x_90 = lean_array_fget(x_44, x_89); -x_91 = lean_unsigned_to_nat(3u); -x_92 = lean_array_fget(x_44, x_91); -lean_dec(x_44); -x_93 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__21; +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_91 = lean_unsigned_to_nat(2u); +x_92 = lean_array_fget(x_46, x_91); +x_93 = lean_unsigned_to_nat(3u); +x_94 = lean_array_fget(x_46, x_93); +lean_dec(x_46); +x_95 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__21; +lean_inc(x_94); lean_inc(x_92); -lean_inc(x_90); -x_94 = l_Lean_mkAppB(x_93, x_90, x_92); -x_95 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__24; -x_96 = l_Lean_mkAppB(x_95, x_90, x_92); -x_97 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_98 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_97, x_94); -x_99 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_98, x_96); -x_100 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_100, 0, x_99); -lean_ctor_set(x_100, 1, x_9); -return x_100; +x_96 = l_Lean_mkAppB(x_95, x_92, x_94); +x_97 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__24; +x_98 = l_Lean_mkAppB(x_97, x_92, x_94); +x_99 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_100 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_99, x_96); +x_101 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_100, x_98); +x_102 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_102, 0, x_101); +lean_ctor_set(x_102, 1, x_11); +return x_102; } } } } else { -lean_object* x_101; uint8_t x_102; -lean_dec(x_46); -x_101 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__25; -x_102 = lean_string_dec_eq(x_45, x_101); -lean_dec(x_45); -if (x_102 == 0) +lean_object* x_103; uint8_t x_104; +lean_dec(x_48); +x_103 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__25; +x_104 = lean_string_dec_eq(x_47, x_103); +lean_dec(x_47); +if (x_104 == 0) { -lean_object* x_103; lean_object* x_104; -lean_dec(x_44); +lean_object* x_105; lean_object* x_106; +lean_dec(x_46); +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_103 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_104 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_104, 0, x_103); -lean_ctor_set(x_104, 1, x_9); -return x_104; +x_105 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_106 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_106, 0, x_105); +lean_ctor_set(x_106, 1, x_11); +return x_106; } else { -lean_object* x_105; lean_object* x_106; uint8_t x_107; -x_105 = lean_array_get_size(x_44); -x_106 = lean_unsigned_to_nat(6u); -x_107 = lean_nat_dec_eq(x_105, x_106); -lean_dec(x_105); -if (x_107 == 0) +lean_object* x_107; lean_object* x_108; uint8_t x_109; +x_107 = lean_array_get_size(x_46); +x_108 = lean_unsigned_to_nat(6u); +x_109 = lean_nat_dec_eq(x_107, x_108); +lean_dec(x_107); +if (x_109 == 0) { -lean_object* x_108; lean_object* x_109; -lean_dec(x_44); +lean_object* x_110; lean_object* x_111; +lean_dec(x_46); +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_108 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_109 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_109, 0, x_108); -lean_ctor_set(x_109, 1, x_9); -return x_109; +x_110 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_111 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_111, 0, x_110); +lean_ctor_set(x_111, 1, x_11); +return x_111; } 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; -x_110 = lean_unsigned_to_nat(4u); -x_111 = lean_array_fget(x_44, x_110); -x_112 = lean_unsigned_to_nat(5u); -x_113 = lean_array_fget(x_44, x_112); -lean_dec(x_44); -lean_inc(x_113); -x_114 = l_Lean_Expr_getAppFnArgs(x_113); -x_115 = lean_ctor_get(x_114, 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; +x_112 = lean_unsigned_to_nat(4u); +x_113 = lean_array_fget(x_46, x_112); +x_114 = lean_unsigned_to_nat(5u); +x_115 = lean_array_fget(x_46, x_114); +lean_dec(x_46); lean_inc(x_115); -if (lean_obj_tag(x_115) == 1) -{ -lean_object* x_116; -x_116 = lean_ctor_get(x_115, 0); -lean_inc(x_116); -if (lean_obj_tag(x_116) == 1) -{ -lean_object* x_117; +x_116 = l_Lean_Expr_getAppFnArgs(x_115); x_117 = lean_ctor_get(x_116, 0); lean_inc(x_117); -if (lean_obj_tag(x_117) == 0) +if (lean_obj_tag(x_117) == 1) { -lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; uint8_t x_122; -x_118 = lean_ctor_get(x_114, 1); +lean_object* x_118; +x_118 = lean_ctor_get(x_117, 0); lean_inc(x_118); -lean_dec(x_114); -x_119 = lean_ctor_get(x_115, 1); +if (lean_obj_tag(x_118) == 1) +{ +lean_object* x_119; +x_119 = lean_ctor_get(x_118, 0); lean_inc(x_119); -lean_dec(x_115); +if (lean_obj_tag(x_119) == 0) +{ +lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; uint8_t x_124; x_120 = lean_ctor_get(x_116, 1); lean_inc(x_120); lean_dec(x_116); -x_121 = l_Lean_Elab_Tactic_Omega_groundNat_x3f___closed__5; -x_122 = lean_string_dec_eq(x_120, x_121); -if (x_122 == 0) +x_121 = lean_ctor_get(x_117, 1); +lean_inc(x_121); +lean_dec(x_117); +x_122 = lean_ctor_get(x_118, 1); +lean_inc(x_122); +lean_dec(x_118); +x_123 = l_Lean_Elab_Tactic_Omega_groundNat_x3f___closed__5; +x_124 = lean_string_dec_eq(x_122, x_123); +if (x_124 == 0) { -uint8_t x_123; -x_123 = lean_string_dec_eq(x_120, x_47); -lean_dec(x_120); -if (x_123 == 0) +uint8_t x_125; +x_125 = lean_string_dec_eq(x_122, x_49); +lean_dec(x_122); +if (x_125 == 0) { -lean_object* x_124; lean_object* x_125; -lean_dec(x_119); -lean_dec(x_118); +lean_object* x_126; lean_object* x_127; +lean_dec(x_121); +lean_dec(x_120); +lean_dec(x_115); lean_dec(x_113); -lean_dec(x_111); +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_124 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_125 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_125, 0, x_124); -lean_ctor_set(x_125, 1, x_9); -return x_125; +x_126 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_127 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_127, 0, x_126); +lean_ctor_set(x_127, 1, x_11); +return x_127; } else { -lean_object* x_126; uint8_t x_127; -x_126 = l_Lean_Elab_Tactic_Omega_natCast_x3f___closed__2; -x_127 = lean_string_dec_eq(x_119, x_126); -lean_dec(x_119); -if (x_127 == 0) +lean_object* x_128; uint8_t x_129; +x_128 = l_Lean_Elab_Tactic_Omega_natCast_x3f___closed__2; +x_129 = lean_string_dec_eq(x_121, x_128); +lean_dec(x_121); +if (x_129 == 0) { -lean_object* x_128; lean_object* x_129; -lean_dec(x_118); +lean_object* x_130; lean_object* x_131; +lean_dec(x_120); +lean_dec(x_115); lean_dec(x_113); -lean_dec(x_111); +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_128 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_129 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_129, 0, x_128); -lean_ctor_set(x_129, 1, x_9); -return x_129; +x_130 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_131 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_131, 0, x_130); +lean_ctor_set(x_131, 1, x_11); +return x_131; } else { -lean_object* x_130; lean_object* x_131; uint8_t x_132; -x_130 = lean_array_get_size(x_118); -x_131 = lean_unsigned_to_nat(3u); -x_132 = lean_nat_dec_eq(x_130, x_131); -lean_dec(x_130); -if (x_132 == 0) +lean_object* x_132; lean_object* x_133; uint8_t x_134; +x_132 = lean_array_get_size(x_120); +x_133 = lean_unsigned_to_nat(3u); +x_134 = lean_nat_dec_eq(x_132, x_133); +lean_dec(x_132); +if (x_134 == 0) { -lean_object* x_133; lean_object* x_134; -lean_dec(x_118); +lean_object* x_135; lean_object* x_136; +lean_dec(x_120); +lean_dec(x_115); lean_dec(x_113); -lean_dec(x_111); +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_133 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_134 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_134, 0, x_133); -lean_ctor_set(x_134, 1, x_9); -return x_134; +x_135 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_136 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_136, 0, x_135); +lean_ctor_set(x_136, 1, x_11); +return x_136; } else { -lean_object* x_135; lean_object* x_136; -x_135 = lean_unsigned_to_nat(0u); -x_136 = lean_array_fget(x_118, x_135); -if (lean_obj_tag(x_136) == 4) -{ -lean_object* x_137; -x_137 = lean_ctor_get(x_136, 0); -lean_inc(x_137); -if (lean_obj_tag(x_137) == 1) -{ -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_139; lean_object* x_140; lean_object* x_141; uint8_t x_142; -x_139 = lean_ctor_get(x_136, 1); +lean_object* x_137; lean_object* x_138; +x_137 = lean_unsigned_to_nat(0u); +x_138 = lean_array_fget(x_120, x_137); +if (lean_obj_tag(x_138) == 4) +{ +lean_object* x_139; +x_139 = lean_ctor_get(x_138, 0); lean_inc(x_139); -lean_dec(x_136); -x_140 = lean_ctor_get(x_137, 1); +if (lean_obj_tag(x_139) == 1) +{ +lean_object* x_140; +x_140 = lean_ctor_get(x_139, 0); lean_inc(x_140); -lean_dec(x_137); -x_141 = l_Lean_Elab_Tactic_Omega_atomsList___rarg___closed__1; -x_142 = lean_string_dec_eq(x_140, x_141); -lean_dec(x_140); -if (x_142 == 0) +if (lean_obj_tag(x_140) == 0) { -lean_object* x_143; lean_object* x_144; +lean_object* x_141; lean_object* x_142; lean_object* x_143; uint8_t x_144; +x_141 = lean_ctor_get(x_138, 1); +lean_inc(x_141); +lean_dec(x_138); +x_142 = lean_ctor_get(x_139, 1); +lean_inc(x_142); lean_dec(x_139); -lean_dec(x_118); +x_143 = l_Lean_Elab_Tactic_Omega_atomsList___rarg___closed__1; +x_144 = lean_string_dec_eq(x_142, x_143); +lean_dec(x_142); +if (x_144 == 0) +{ +lean_object* x_145; lean_object* x_146; +lean_dec(x_141); +lean_dec(x_120); +lean_dec(x_115); lean_dec(x_113); -lean_dec(x_111); +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_143 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_144 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_144, 0, x_143); -lean_ctor_set(x_144, 1, x_9); -return x_144; +x_145 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_146 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_146, 0, x_145); +lean_ctor_set(x_146, 1, x_11); +return x_146; } else { -if (lean_obj_tag(x_139) == 0) +if (lean_obj_tag(x_141) == 0) { -lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; -x_145 = lean_unsigned_to_nat(2u); -x_146 = lean_array_fget(x_118, x_145); -lean_dec(x_118); -lean_inc(x_146); -x_147 = l_Lean_Expr_getAppFnArgs(x_146); -x_148 = lean_ctor_get(x_147, 0); +lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; +x_147 = lean_unsigned_to_nat(2u); +x_148 = lean_array_fget(x_120, x_147); +lean_dec(x_120); lean_inc(x_148); -if (lean_obj_tag(x_148) == 1) -{ -lean_object* x_149; -x_149 = lean_ctor_get(x_148, 0); -lean_inc(x_149); -if (lean_obj_tag(x_149) == 1) -{ -lean_object* x_150; +x_149 = l_Lean_Expr_getAppFnArgs(x_148); x_150 = lean_ctor_get(x_149, 0); lean_inc(x_150); -if (lean_obj_tag(x_150) == 0) +if (lean_obj_tag(x_150) == 1) { -lean_object* x_151; lean_object* x_152; lean_object* x_153; uint8_t x_154; -x_151 = lean_ctor_get(x_147, 1); +lean_object* x_151; +x_151 = lean_ctor_get(x_150, 0); lean_inc(x_151); -lean_dec(x_147); -x_152 = lean_ctor_get(x_148, 1); +if (lean_obj_tag(x_151) == 1) +{ +lean_object* x_152; +x_152 = lean_ctor_get(x_151, 0); lean_inc(x_152); -lean_dec(x_148); +if (lean_obj_tag(x_152) == 0) +{ +lean_object* x_153; lean_object* x_154; lean_object* x_155; uint8_t x_156; x_153 = lean_ctor_get(x_149, 1); lean_inc(x_153); lean_dec(x_149); -x_154 = lean_string_dec_eq(x_153, x_121); -lean_dec(x_153); -if (x_154 == 0) -{ -lean_object* x_155; lean_object* x_156; -lean_dec(x_152); +x_154 = lean_ctor_get(x_150, 1); +lean_inc(x_154); +lean_dec(x_150); +x_155 = lean_ctor_get(x_151, 1); +lean_inc(x_155); lean_dec(x_151); -lean_dec(x_146); +x_156 = lean_string_dec_eq(x_155, x_123); +lean_dec(x_155); +if (x_156 == 0) +{ +lean_object* x_157; lean_object* x_158; +lean_dec(x_154); +lean_dec(x_153); +lean_dec(x_148); +lean_dec(x_115); lean_dec(x_113); -lean_dec(x_111); +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_155 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_156 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_156, 0, x_155); -lean_ctor_set(x_156, 1, x_9); -return x_156; +x_157 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_158 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_158, 0, x_157); +lean_ctor_set(x_158, 1, x_11); +return x_158; } else { -lean_object* x_157; uint8_t x_158; -x_157 = l_Lean_Elab_Tactic_Omega_groundNat_x3f___closed__6; -x_158 = lean_string_dec_eq(x_152, x_157); -lean_dec(x_152); -if (x_158 == 0) +lean_object* x_159; uint8_t x_160; +x_159 = l_Lean_Elab_Tactic_Omega_groundNat_x3f___closed__6; +x_160 = lean_string_dec_eq(x_154, x_159); +lean_dec(x_154); +if (x_160 == 0) { -lean_object* x_159; lean_object* x_160; -lean_dec(x_151); -lean_dec(x_146); +lean_object* x_161; lean_object* x_162; +lean_dec(x_153); +lean_dec(x_148); +lean_dec(x_115); lean_dec(x_113); -lean_dec(x_111); +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_159 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_160 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_160, 0, x_159); -lean_ctor_set(x_160, 1, x_9); -return x_160; +x_161 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_162 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_162, 0, x_161); +lean_ctor_set(x_162, 1, x_11); +return x_162; } else { -lean_object* x_161; uint8_t x_162; -x_161 = lean_array_get_size(x_151); -x_162 = lean_nat_dec_eq(x_161, x_106); -lean_dec(x_161); -if (x_162 == 0) +lean_object* x_163; uint8_t x_164; +x_163 = lean_array_get_size(x_153); +x_164 = lean_nat_dec_eq(x_163, x_108); +lean_dec(x_163); +if (x_164 == 0) { -lean_object* x_163; lean_object* x_164; -lean_dec(x_151); -lean_dec(x_146); +lean_object* x_165; lean_object* x_166; +lean_dec(x_153); +lean_dec(x_148); +lean_dec(x_115); lean_dec(x_113); -lean_dec(x_111); +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_163 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_164 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_164, 0, x_163); -lean_ctor_set(x_164, 1, x_9); -return x_164; +x_165 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_166 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_166, 0, x_165); +lean_ctor_set(x_166, 1, x_11); +return x_166; } else { -lean_object* x_165; lean_object* x_166; lean_object* x_167; -x_165 = lean_array_fget(x_151, x_110); -x_166 = lean_array_fget(x_151, x_112); -lean_dec(x_151); -lean_inc(x_165); -x_167 = l_Lean_Elab_Tactic_Omega_natCast_x3f(x_165); -if (lean_obj_tag(x_167) == 0) -{ -lean_object* x_168; lean_object* x_169; -lean_dec(x_166); -lean_dec(x_165); -lean_dec(x_146); +lean_object* x_167; lean_object* x_168; lean_object* x_169; +x_167 = lean_array_fget(x_153, x_112); +x_168 = lean_array_fget(x_153, x_114); +lean_dec(x_153); +lean_inc(x_167); +x_169 = l_Lean_Elab_Tactic_Omega_natCast_x3f(x_167); +if (lean_obj_tag(x_169) == 0) +{ +lean_object* x_170; lean_object* x_171; +lean_dec(x_168); +lean_dec(x_167); +lean_dec(x_148); +lean_dec(x_115); lean_dec(x_113); -lean_dec(x_111); +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_168 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_169 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_169, 0, x_168); -lean_ctor_set(x_169, 1, x_9); -return x_169; -} -else -{ -lean_object* x_170; uint8_t x_171; -x_170 = lean_ctor_get(x_167, 0); -lean_inc(x_170); -lean_dec(x_167); -x_171 = lean_nat_dec_eq(x_170, x_135); -lean_dec(x_170); -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; -x_172 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__30; -x_173 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__32; -x_174 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__35; -x_175 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__36; -lean_inc(x_165); -x_176 = l_Lean_mkApp4(x_172, x_173, x_174, x_175, x_165); -x_177 = l_Lean_Meta_mkDecideProof(x_176, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_177) == 0) -{ -uint8_t x_178; -x_178 = !lean_is_exclusive(x_177); -if (x_178 == 0) -{ -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; uint8_t x_186; -x_179 = lean_ctor_get(x_177, 0); -x_180 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__39; -x_181 = l_Lean_mkApp3(x_180, x_165, x_166, x_179); -x_182 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__42; -x_183 = l_Lean_mkAppB(x_182, x_146, x_181); -x_184 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__53; -lean_inc(x_183); -lean_inc(x_113); -lean_inc(x_111); -x_185 = l_Lean_mkApp3(x_184, x_111, x_113, x_183); -x_186 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__50; -if (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; lean_object* x_193; lean_object* x_194; -x_187 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__48; -x_188 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__65; +x_170 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_171 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_171, 0, x_170); +lean_ctor_set(x_171, 1, x_11); +return x_171; +} +else +{ +lean_object* x_172; uint8_t x_173; +x_172 = lean_ctor_get(x_169, 0); +lean_inc(x_172); +lean_dec(x_169); +x_173 = lean_nat_dec_eq(x_172, x_137); +lean_dec(x_172); +if (x_173 == 0) +{ +lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; +x_174 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__30; +x_175 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__32; +x_176 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__35; +x_177 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__36; +lean_inc(x_167); +x_178 = l_Lean_mkApp4(x_174, x_175, x_176, x_177, x_167); +x_179 = l_Lean_Meta_mkDecideProof(x_178, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_179) == 0) +{ +uint8_t x_180; +x_180 = !lean_is_exclusive(x_179); +if (x_180 == 0) +{ +lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; uint8_t x_188; +x_181 = lean_ctor_get(x_179, 0); +x_182 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__39; +x_183 = l_Lean_mkApp3(x_182, x_167, x_168, x_181); +x_184 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__42; +x_185 = l_Lean_mkAppB(x_184, x_148, x_183); +x_186 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__53; +lean_inc(x_185); +lean_inc(x_115); lean_inc(x_113); -x_189 = l_Lean_mkApp3(x_187, x_113, x_188, x_183); -x_190 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__45; -x_191 = l_Lean_mkApp3(x_190, x_111, x_113, x_189); -x_192 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_193 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_192, x_191); -x_194 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_193, x_185); -lean_ctor_set(x_177, 0, x_194); -return x_177; +x_187 = l_Lean_mkApp3(x_186, x_113, x_115, x_185); +x_188 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__50; +if (x_188 == 0) +{ +lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; +x_189 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__48; +x_190 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__65; +lean_inc(x_115); +x_191 = l_Lean_mkApp3(x_189, x_115, x_190, x_185); +x_192 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__45; +x_193 = l_Lean_mkApp3(x_192, x_113, x_115, x_191); +x_194 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_195 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_194, x_193); +x_196 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_195, x_187); +lean_ctor_set(x_179, 0, x_196); +return x_179; } else { -lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; -x_195 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__48; -x_196 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__67; -lean_inc(x_113); -x_197 = l_Lean_mkApp3(x_195, x_113, x_196, x_183); -x_198 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__45; -x_199 = l_Lean_mkApp3(x_198, x_111, x_113, x_197); -x_200 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_201 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_200, x_199); -x_202 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_201, x_185); -lean_ctor_set(x_177, 0, x_202); -return x_177; -} -} -else -{ -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; uint8_t x_211; -x_203 = lean_ctor_get(x_177, 0); -x_204 = lean_ctor_get(x_177, 1); -lean_inc(x_204); -lean_inc(x_203); -lean_dec(x_177); -x_205 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__39; -x_206 = l_Lean_mkApp3(x_205, x_165, x_166, x_203); -x_207 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__42; -x_208 = l_Lean_mkAppB(x_207, x_146, x_206); -x_209 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__53; -lean_inc(x_208); -lean_inc(x_113); -lean_inc(x_111); -x_210 = l_Lean_mkApp3(x_209, x_111, x_113, x_208); -x_211 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__50; -if (x_211 == 0) -{ -lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; -x_212 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__48; -x_213 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__65; -lean_inc(x_113); -x_214 = l_Lean_mkApp3(x_212, x_113, x_213, x_208); -x_215 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__45; -x_216 = l_Lean_mkApp3(x_215, x_111, x_113, x_214); -x_217 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_218 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_217, x_216); -x_219 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_218, x_210); -x_220 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_220, 0, x_219); -lean_ctor_set(x_220, 1, x_204); -return x_220; -} -else -{ -lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; -x_221 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__48; -x_222 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__67; +lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; +x_197 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__48; +x_198 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__67; +lean_inc(x_115); +x_199 = l_Lean_mkApp3(x_197, x_115, x_198, x_185); +x_200 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__45; +x_201 = l_Lean_mkApp3(x_200, x_113, x_115, x_199); +x_202 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_203 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_202, x_201); +x_204 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_203, x_187); +lean_ctor_set(x_179, 0, x_204); +return x_179; +} +} +else +{ +lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; uint8_t x_213; +x_205 = lean_ctor_get(x_179, 0); +x_206 = lean_ctor_get(x_179, 1); +lean_inc(x_206); +lean_inc(x_205); +lean_dec(x_179); +x_207 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__39; +x_208 = l_Lean_mkApp3(x_207, x_167, x_168, x_205); +x_209 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__42; +x_210 = l_Lean_mkAppB(x_209, x_148, x_208); +x_211 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__53; +lean_inc(x_210); +lean_inc(x_115); lean_inc(x_113); -x_223 = l_Lean_mkApp3(x_221, x_113, x_222, x_208); -x_224 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__45; -x_225 = l_Lean_mkApp3(x_224, x_111, x_113, x_223); -x_226 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_227 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_226, x_225); -x_228 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_227, x_210); -x_229 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_229, 0, x_228); -lean_ctor_set(x_229, 1, x_204); -return x_229; +x_212 = l_Lean_mkApp3(x_211, x_113, x_115, x_210); +x_213 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__50; +if (x_213 == 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; +x_214 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__48; +x_215 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__65; +lean_inc(x_115); +x_216 = l_Lean_mkApp3(x_214, x_115, x_215, x_210); +x_217 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__45; +x_218 = l_Lean_mkApp3(x_217, x_113, x_115, x_216); +x_219 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_220 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_219, x_218); +x_221 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_220, x_212); +x_222 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_222, 0, x_221); +lean_ctor_set(x_222, 1, x_206); +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 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__48; +x_224 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__67; +lean_inc(x_115); +x_225 = l_Lean_mkApp3(x_223, x_115, x_224, x_210); +x_226 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__45; +x_227 = l_Lean_mkApp3(x_226, x_113, x_115, x_225); +x_228 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_229 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_228, x_227); +x_230 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_229, x_212); +x_231 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_231, 0, x_230); +lean_ctor_set(x_231, 1, x_206); +return x_231; } } } else { -uint8_t x_230; -lean_dec(x_166); -lean_dec(x_165); -lean_dec(x_146); +uint8_t x_232; +lean_dec(x_168); +lean_dec(x_167); +lean_dec(x_148); +lean_dec(x_115); lean_dec(x_113); -lean_dec(x_111); -x_230 = !lean_is_exclusive(x_177); -if (x_230 == 0) +x_232 = !lean_is_exclusive(x_179); +if (x_232 == 0) { -return x_177; -} -else -{ -lean_object* x_231; lean_object* x_232; lean_object* x_233; -x_231 = lean_ctor_get(x_177, 0); -x_232 = lean_ctor_get(x_177, 1); -lean_inc(x_232); -lean_inc(x_231); -lean_dec(x_177); -x_233 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_233, 0, x_231); -lean_ctor_set(x_233, 1, x_232); -return x_233; -} -} +return x_179; } else { -lean_object* x_234; lean_object* x_235; -lean_dec(x_166); -lean_dec(x_165); -lean_dec(x_146); -lean_dec(x_113); -lean_dec(x_111); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -x_234 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_235 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_235, 0, x_234); -lean_ctor_set(x_235, 1, x_9); +lean_object* x_233; lean_object* x_234; lean_object* x_235; +x_233 = lean_ctor_get(x_179, 0); +x_234 = lean_ctor_get(x_179, 1); +lean_inc(x_234); +lean_inc(x_233); +lean_dec(x_179); +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; } } } -} -} -} else { lean_object* x_236; lean_object* x_237; -lean_dec(x_150); -lean_dec(x_149); +lean_dec(x_168); +lean_dec(x_167); lean_dec(x_148); -lean_dec(x_147); -lean_dec(x_146); +lean_dec(x_115); lean_dec(x_113); -lean_dec(x_111); +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_236 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; x_237 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_237, 0, x_236); -lean_ctor_set(x_237, 1, x_9); +lean_ctor_set(x_237, 1, x_11); return x_237; } } +} +} +} +} else { lean_object* x_238; lean_object* x_239; +lean_dec(x_152); +lean_dec(x_151); +lean_dec(x_150); lean_dec(x_149); lean_dec(x_148); -lean_dec(x_147); -lean_dec(x_146); +lean_dec(x_115); lean_dec(x_113); -lean_dec(x_111); +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_238 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; x_239 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_239, 0, x_238); -lean_ctor_set(x_239, 1, x_9); +lean_ctor_set(x_239, 1, x_11); return x_239; } } else { lean_object* x_240; lean_object* x_241; +lean_dec(x_151); +lean_dec(x_150); +lean_dec(x_149); lean_dec(x_148); -lean_dec(x_147); -lean_dec(x_146); +lean_dec(x_115); lean_dec(x_113); -lean_dec(x_111); +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_240 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; x_241 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_241, 0, x_240); -lean_ctor_set(x_241, 1, x_9); +lean_ctor_set(x_241, 1, x_11); return x_241; } } else { lean_object* x_242; lean_object* x_243; -lean_dec(x_139); -lean_dec(x_118); +lean_dec(x_150); +lean_dec(x_149); +lean_dec(x_148); +lean_dec(x_115); lean_dec(x_113); -lean_dec(x_111); +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_242 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; x_243 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_243, 0, x_242); -lean_ctor_set(x_243, 1, x_9); +lean_ctor_set(x_243, 1, x_11); return x_243; } } -} else { lean_object* x_244; lean_object* x_245; -lean_dec(x_138); -lean_dec(x_137); -lean_dec(x_136); -lean_dec(x_118); +lean_dec(x_141); +lean_dec(x_120); +lean_dec(x_115); lean_dec(x_113); -lean_dec(x_111); +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_244 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; x_245 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_245, 0, x_244); -lean_ctor_set(x_245, 1, x_9); +lean_ctor_set(x_245, 1, x_11); return x_245; } } +} else { lean_object* x_246; lean_object* x_247; -lean_dec(x_137); -lean_dec(x_136); -lean_dec(x_118); +lean_dec(x_140); +lean_dec(x_139); +lean_dec(x_138); +lean_dec(x_120); +lean_dec(x_115); lean_dec(x_113); -lean_dec(x_111); +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_246 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; x_247 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_247, 0, x_246); -lean_ctor_set(x_247, 1, x_9); +lean_ctor_set(x_247, 1, x_11); return x_247; } } else { lean_object* x_248; lean_object* x_249; -lean_dec(x_136); -lean_dec(x_118); +lean_dec(x_139); +lean_dec(x_138); +lean_dec(x_120); +lean_dec(x_115); lean_dec(x_113); -lean_dec(x_111); +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_248 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; x_249 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_249, 0, x_248); -lean_ctor_set(x_249, 1, x_9); +lean_ctor_set(x_249, 1, x_11); return x_249; } } -} -} -} else { -lean_object* x_250; uint8_t x_251; +lean_object* x_250; lean_object* x_251; +lean_dec(x_138); lean_dec(x_120); -x_250 = l_Lean_Elab_Tactic_Omega_groundNat_x3f___closed__6; -x_251 = lean_string_dec_eq(x_119, x_250); -lean_dec(x_119); -if (x_251 == 0) -{ -lean_object* x_252; lean_object* x_253; -lean_dec(x_118); +lean_dec(x_115); lean_dec(x_113); -lean_dec(x_111); +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_252 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_253 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_253, 0, x_252); -lean_ctor_set(x_253, 1, x_9); -return x_253; +x_250 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_251 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_251, 0, x_250); +lean_ctor_set(x_251, 1, x_11); +return x_251; +} +} +} +} } else { -lean_object* x_254; uint8_t x_255; -x_254 = lean_array_get_size(x_118); -x_255 = lean_nat_dec_eq(x_254, x_106); -lean_dec(x_254); -if (x_255 == 0) +lean_object* x_252; uint8_t x_253; +lean_dec(x_122); +x_252 = l_Lean_Elab_Tactic_Omega_groundNat_x3f___closed__6; +x_253 = lean_string_dec_eq(x_121, x_252); +lean_dec(x_121); +if (x_253 == 0) { -lean_object* x_256; lean_object* x_257; -lean_dec(x_118); +lean_object* x_254; lean_object* x_255; +lean_dec(x_120); +lean_dec(x_115); lean_dec(x_113); -lean_dec(x_111); +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_256 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_257 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_257, 0, x_256); -lean_ctor_set(x_257, 1, x_9); -return x_257; +x_254 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_255 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_255, 0, x_254); +lean_ctor_set(x_255, 1, x_11); +return x_255; } else { -lean_object* x_258; lean_object* x_259; lean_object* x_260; -x_258 = lean_array_fget(x_118, x_110); -x_259 = lean_array_fget(x_118, x_112); -lean_dec(x_118); -lean_inc(x_258); -x_260 = l_Lean_Elab_Tactic_Omega_natCast_x3f(x_258); -if (lean_obj_tag(x_260) == 0) +lean_object* x_256; uint8_t x_257; +x_256 = lean_array_get_size(x_120); +x_257 = lean_nat_dec_eq(x_256, x_108); +lean_dec(x_256); +if (x_257 == 0) { -lean_object* x_261; lean_object* x_262; -lean_dec(x_259); -lean_dec(x_258); +lean_object* x_258; lean_object* x_259; +lean_dec(x_120); +lean_dec(x_115); lean_dec(x_113); -lean_dec(x_111); +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_261 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_262 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_262, 0, x_261); -lean_ctor_set(x_262, 1, x_9); -return x_262; +x_258 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_259 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_259, 0, x_258); +lean_ctor_set(x_259, 1, x_11); +return x_259; } else { -lean_object* x_263; lean_object* x_264; uint8_t x_265; -x_263 = lean_ctor_get(x_260, 0); -lean_inc(x_263); -lean_dec(x_260); -x_264 = lean_unsigned_to_nat(0u); -x_265 = lean_nat_dec_eq(x_263, x_264); -lean_dec(x_263); -if (x_265 == 0) -{ -lean_object* x_266; uint8_t x_304; -x_304 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__50; -if (x_304 == 0) -{ -lean_object* x_305; -x_305 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__73; -x_266 = x_305; -goto block_303; -} -else -{ -lean_object* x_306; -x_306 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__67; -x_266 = x_306; -goto block_303; -} -block_303: -{ -lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; -x_267 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__30; -x_268 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__58; -x_269 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__70; -lean_inc(x_258); -lean_inc(x_266); -x_270 = l_Lean_mkApp4(x_267, x_268, x_269, x_266, x_258); -x_271 = l_Lean_Meta_mkDecideProof(x_270, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_271) == 0) -{ -uint8_t x_272; -x_272 = !lean_is_exclusive(x_271); -if (x_272 == 0) -{ -lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; -x_273 = lean_ctor_get(x_271, 0); -x_274 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__72; -x_275 = l_Lean_mkApp3(x_274, x_258, x_259, x_273); -x_276 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__48; -lean_inc(x_275); -lean_inc(x_113); -x_277 = l_Lean_mkApp3(x_276, x_113, x_266, x_275); -x_278 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__45; -lean_inc(x_113); -lean_inc(x_111); -x_279 = l_Lean_mkApp3(x_278, x_111, x_113, x_277); -x_280 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__53; -x_281 = l_Lean_mkApp3(x_280, x_111, x_113, x_275); -x_282 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_283 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_282, x_279); -x_284 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_283, x_281); -lean_ctor_set(x_271, 0, x_284); -return x_271; -} -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_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; -x_285 = lean_ctor_get(x_271, 0); -x_286 = lean_ctor_get(x_271, 1); -lean_inc(x_286); -lean_inc(x_285); -lean_dec(x_271); -x_287 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__72; -x_288 = l_Lean_mkApp3(x_287, x_258, x_259, x_285); -x_289 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__48; -lean_inc(x_288); -lean_inc(x_113); -x_290 = l_Lean_mkApp3(x_289, x_113, x_266, x_288); -x_291 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__45; -lean_inc(x_113); -lean_inc(x_111); -x_292 = l_Lean_mkApp3(x_291, x_111, x_113, x_290); -x_293 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__53; -x_294 = l_Lean_mkApp3(x_293, x_111, x_113, x_288); -x_295 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_296 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_295, x_292); -x_297 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_296, x_294); -x_298 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_298, 0, x_297); -lean_ctor_set(x_298, 1, x_286); -return x_298; -} -} -else +lean_object* x_260; lean_object* x_261; lean_object* x_262; +x_260 = lean_array_fget(x_120, x_112); +x_261 = lean_array_fget(x_120, x_114); +lean_dec(x_120); +lean_inc(x_260); +x_262 = l_Lean_Elab_Tactic_Omega_natCast_x3f(x_260); +if (lean_obj_tag(x_262) == 0) { -uint8_t x_299; -lean_dec(x_266); -lean_dec(x_259); -lean_dec(x_258); +lean_object* x_263; lean_object* x_264; +lean_dec(x_261); +lean_dec(x_260); +lean_dec(x_115); lean_dec(x_113); -lean_dec(x_111); -x_299 = !lean_is_exclusive(x_271); -if (x_299 == 0) -{ -return x_271; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +x_263 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_264 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_264, 0, x_263); +lean_ctor_set(x_264, 1, x_11); +return x_264; +} +else +{ +lean_object* x_265; lean_object* x_266; uint8_t x_267; +x_265 = lean_ctor_get(x_262, 0); +lean_inc(x_265); +lean_dec(x_262); +x_266 = lean_unsigned_to_nat(0u); +x_267 = lean_nat_dec_eq(x_265, x_266); +lean_dec(x_265); +if (x_267 == 0) +{ +lean_object* x_268; uint8_t x_306; +x_306 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__50; +if (x_306 == 0) +{ +lean_object* x_307; +x_307 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__73; +x_268 = x_307; +goto block_305; +} +else +{ +lean_object* x_308; +x_308 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__67; +x_268 = x_308; +goto block_305; +} +block_305: +{ +lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; +x_269 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__30; +x_270 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__58; +x_271 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__70; +lean_inc(x_260); +lean_inc(x_268); +x_272 = l_Lean_mkApp4(x_269, x_270, x_271, x_268, x_260); +x_273 = l_Lean_Meta_mkDecideProof(x_272, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_273) == 0) +{ +uint8_t x_274; +x_274 = !lean_is_exclusive(x_273); +if (x_274 == 0) +{ +lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; +x_275 = lean_ctor_get(x_273, 0); +x_276 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__72; +x_277 = l_Lean_mkApp3(x_276, x_260, x_261, x_275); +x_278 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__48; +lean_inc(x_277); +lean_inc(x_115); +x_279 = l_Lean_mkApp3(x_278, x_115, x_268, x_277); +x_280 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__45; +lean_inc(x_115); +lean_inc(x_113); +x_281 = l_Lean_mkApp3(x_280, x_113, x_115, x_279); +x_282 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__53; +x_283 = l_Lean_mkApp3(x_282, x_113, x_115, x_277); +x_284 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_285 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_284, x_281); +x_286 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_285, x_283); +lean_ctor_set(x_273, 0, x_286); +return x_273; } else { -lean_object* x_300; lean_object* x_301; lean_object* x_302; -x_300 = lean_ctor_get(x_271, 0); -x_301 = lean_ctor_get(x_271, 1); -lean_inc(x_301); -lean_inc(x_300); -lean_dec(x_271); -x_302 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_302, 0, x_300); -lean_ctor_set(x_302, 1, x_301); -return x_302; -} -} +lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; +x_287 = lean_ctor_get(x_273, 0); +x_288 = lean_ctor_get(x_273, 1); +lean_inc(x_288); +lean_inc(x_287); +lean_dec(x_273); +x_289 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__72; +x_290 = l_Lean_mkApp3(x_289, x_260, x_261, x_287); +x_291 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__48; +lean_inc(x_290); +lean_inc(x_115); +x_292 = l_Lean_mkApp3(x_291, x_115, x_268, x_290); +x_293 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__45; +lean_inc(x_115); +lean_inc(x_113); +x_294 = l_Lean_mkApp3(x_293, x_113, x_115, x_292); +x_295 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__53; +x_296 = l_Lean_mkApp3(x_295, x_113, x_115, x_290); +x_297 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_298 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_297, x_294); +x_299 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_298, x_296); +x_300 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_300, 0, x_299); +lean_ctor_set(x_300, 1, x_288); +return x_300; } } else { -lean_object* x_307; lean_object* x_308; -lean_dec(x_259); -lean_dec(x_258); +uint8_t x_301; +lean_dec(x_268); +lean_dec(x_261); +lean_dec(x_260); +lean_dec(x_115); lean_dec(x_113); -lean_dec(x_111); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -x_307 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_308 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_308, 0, x_307); -lean_ctor_set(x_308, 1, x_9); -return x_308; -} +x_301 = !lean_is_exclusive(x_273); +if (x_301 == 0) +{ +return x_273; } +else +{ +lean_object* x_302; lean_object* x_303; lean_object* x_304; +x_302 = lean_ctor_get(x_273, 0); +x_303 = lean_ctor_get(x_273, 1); +lean_inc(x_303); +lean_inc(x_302); +lean_dec(x_273); +x_304 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_304, 0, x_302); +lean_ctor_set(x_304, 1, x_303); +return x_304; } } } @@ -4888,297 +5117,319 @@ return x_308; else { lean_object* x_309; lean_object* x_310; -lean_dec(x_117); -lean_dec(x_116); +lean_dec(x_261); +lean_dec(x_260); lean_dec(x_115); -lean_dec(x_114); lean_dec(x_113); -lean_dec(x_111); +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_309 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; x_310 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_310, 0, x_309); -lean_ctor_set(x_310, 1, x_9); +lean_ctor_set(x_310, 1, x_11); return x_310; } } +} +} +} +} else { lean_object* x_311; lean_object* x_312; +lean_dec(x_119); +lean_dec(x_118); +lean_dec(x_117); lean_dec(x_116); lean_dec(x_115); -lean_dec(x_114); lean_dec(x_113); -lean_dec(x_111); +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_311 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; x_312 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_312, 0, x_311); -lean_ctor_set(x_312, 1, x_9); +lean_ctor_set(x_312, 1, x_11); return x_312; } } else { lean_object* x_313; lean_object* x_314; +lean_dec(x_118); +lean_dec(x_117); +lean_dec(x_116); lean_dec(x_115); -lean_dec(x_114); lean_dec(x_113); -lean_dec(x_111); +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_313 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; x_314 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_314, 0, x_313); -lean_ctor_set(x_314, 1, x_9); +lean_ctor_set(x_314, 1, x_11); return x_314; } } +else +{ +lean_object* x_315; lean_object* x_316; +lean_dec(x_117); +lean_dec(x_116); +lean_dec(x_115); +lean_dec(x_113); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +x_315 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_316 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_316, 0, x_315); +lean_ctor_set(x_316, 1, x_11); +return x_316; +} +} } } } else { -lean_object* x_315; uint8_t x_316; -lean_dec(x_46); -x_315 = l_Lean_Elab_Tactic_Omega_groundNat_x3f___closed__8; -x_316 = lean_string_dec_eq(x_45, x_315); -lean_dec(x_45); -if (x_316 == 0) +lean_object* x_317; uint8_t x_318; +lean_dec(x_48); +x_317 = l_Lean_Elab_Tactic_Omega_groundNat_x3f___closed__8; +x_318 = lean_string_dec_eq(x_47, x_317); +lean_dec(x_47); +if (x_318 == 0) { -lean_object* x_317; lean_object* x_318; -lean_dec(x_44); +lean_object* x_319; lean_object* x_320; +lean_dec(x_46); +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_317 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_318 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_318, 0, x_317); -lean_ctor_set(x_318, 1, x_9); -return x_318; +x_319 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_320 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_320, 0, x_319); +lean_ctor_set(x_320, 1, x_11); +return x_320; } else { -lean_object* x_319; lean_object* x_320; uint8_t x_321; -x_319 = lean_array_get_size(x_44); -x_320 = lean_unsigned_to_nat(6u); -x_321 = lean_nat_dec_eq(x_319, x_320); -lean_dec(x_319); -if (x_321 == 0) +lean_object* x_321; lean_object* x_322; uint8_t x_323; +x_321 = lean_array_get_size(x_46); +x_322 = lean_unsigned_to_nat(6u); +x_323 = lean_nat_dec_eq(x_321, x_322); +lean_dec(x_321); +if (x_323 == 0) { -lean_object* x_322; lean_object* x_323; -lean_dec(x_44); +lean_object* x_324; lean_object* x_325; +lean_dec(x_46); +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_322 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_323 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_323, 0, x_322); -lean_ctor_set(x_323, 1, x_9); -return x_323; +x_324 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_325 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_325, 0, x_324); +lean_ctor_set(x_325, 1, x_11); +return x_325; } else { -lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; -x_324 = lean_unsigned_to_nat(4u); -x_325 = lean_array_fget(x_44, x_324); -x_326 = lean_unsigned_to_nat(5u); -x_327 = lean_array_fget(x_44, x_326); -lean_dec(x_44); -lean_inc(x_327); -x_328 = l_Lean_Elab_Tactic_Omega_natCast_x3f(x_327); -if (lean_obj_tag(x_328) == 0) +lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; +x_326 = lean_unsigned_to_nat(4u); +x_327 = lean_array_fget(x_46, x_326); +x_328 = lean_unsigned_to_nat(5u); +x_329 = lean_array_fget(x_46, x_328); +lean_dec(x_46); +lean_inc(x_329); +x_330 = l_Lean_Elab_Tactic_Omega_natCast_x3f(x_329); +if (lean_obj_tag(x_330) == 0) { -lean_object* x_329; lean_object* x_330; +lean_object* x_331; lean_object* x_332; +lean_dec(x_329); lean_dec(x_327); -lean_dec(x_325); +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_329 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_330 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_330, 0, x_329); -lean_ctor_set(x_330, 1, x_9); -return x_330; -} -else -{ -lean_object* x_331; lean_object* x_332; uint8_t x_333; -x_331 = lean_ctor_get(x_328, 0); -lean_inc(x_331); -lean_dec(x_328); -x_332 = lean_unsigned_to_nat(0u); -x_333 = lean_nat_dec_eq(x_331, x_332); -lean_dec(x_331); -if (x_333 == 0) -{ -lean_object* x_334; uint8_t x_373; -x_373 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__50; -if (x_373 == 0) -{ -lean_object* x_374; -x_374 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__87; -x_334 = x_374; -goto block_372; -} -else -{ -lean_object* x_375; -x_375 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__67; -x_334 = x_375; -goto block_372; -} -block_372: -{ -lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; -x_335 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__78; -x_336 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__58; -lean_inc(x_334); -lean_inc(x_327); -x_337 = l_Lean_mkApp3(x_335, x_336, x_327, x_334); -x_338 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__79; -x_339 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__70; -lean_inc(x_327); -x_340 = l_Lean_mkApp4(x_338, x_336, x_339, x_334, x_327); +x_331 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_332 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_332, 0, x_331); +lean_ctor_set(x_332, 1, x_11); +return x_332; +} +else +{ +lean_object* x_333; lean_object* x_334; uint8_t x_335; +x_333 = lean_ctor_get(x_330, 0); +lean_inc(x_333); +lean_dec(x_330); +x_334 = lean_unsigned_to_nat(0u); +x_335 = lean_nat_dec_eq(x_333, x_334); +lean_dec(x_333); +if (x_335 == 0) +{ +lean_object* x_336; uint8_t x_375; +x_375 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__50; +if (x_375 == 0) +{ +lean_object* x_376; +x_376 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__87; +x_336 = x_376; +goto block_374; +} +else +{ +lean_object* x_377; +x_377 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__67; +x_336 = x_377; +goto block_374; +} +block_374: +{ +lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; +x_337 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__78; +x_338 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__58; +lean_inc(x_336); +lean_inc(x_329); +x_339 = l_Lean_mkApp3(x_337, x_338, x_329, x_336); +x_340 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__79; +x_341 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__70; +lean_inc(x_329); +x_342 = l_Lean_mkApp4(x_340, x_338, x_341, x_336, x_329); +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_341 = l_Lean_Meta_mkDecideProof(x_337, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_341) == 0) -{ -lean_object* x_342; lean_object* x_343; lean_object* x_344; -x_342 = lean_ctor_get(x_341, 0); -lean_inc(x_342); -x_343 = lean_ctor_get(x_341, 1); -lean_inc(x_343); -lean_dec(x_341); -x_344 = l_Lean_Meta_mkDecideProof(x_340, x_5, x_6, x_7, x_8, x_343); -if (lean_obj_tag(x_344) == 0) -{ -uint8_t x_345; -x_345 = !lean_is_exclusive(x_344); -if (x_345 == 0) -{ -lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; -x_346 = lean_ctor_get(x_344, 0); -x_347 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__82; +x_343 = l_Lean_Meta_mkDecideProof(x_339, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_343) == 0) +{ +lean_object* x_344; lean_object* x_345; lean_object* x_346; +x_344 = lean_ctor_get(x_343, 0); +lean_inc(x_344); +x_345 = lean_ctor_get(x_343, 1); +lean_inc(x_345); +lean_dec(x_343); +x_346 = l_Lean_Meta_mkDecideProof(x_342, x_7, x_8, x_9, x_10, x_345); +if (lean_obj_tag(x_346) == 0) +{ +uint8_t x_347; +x_347 = !lean_is_exclusive(x_346); +if (x_347 == 0) +{ +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_348 = lean_ctor_get(x_346, 0); +x_349 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__82; +lean_inc(x_329); lean_inc(x_327); -lean_inc(x_325); -x_348 = l_Lean_mkApp3(x_347, x_325, x_327, x_342); -x_349 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__85; -x_350 = l_Lean_mkApp3(x_349, x_325, x_327, x_346); -x_351 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_352 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_351, x_348); -x_353 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_352, x_350); -lean_ctor_set(x_344, 0, x_353); -return x_344; -} -else -{ -lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; -x_354 = lean_ctor_get(x_344, 0); -x_355 = lean_ctor_get(x_344, 1); -lean_inc(x_355); -lean_inc(x_354); -lean_dec(x_344); -x_356 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__82; +x_350 = l_Lean_mkApp3(x_349, x_327, x_329, x_344); +x_351 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__85; +x_352 = l_Lean_mkApp3(x_351, x_327, x_329, x_348); +x_353 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_354 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_353, x_350); +x_355 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_354, x_352); +lean_ctor_set(x_346, 0, x_355); +return x_346; +} +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; +x_356 = lean_ctor_get(x_346, 0); +x_357 = lean_ctor_get(x_346, 1); +lean_inc(x_357); +lean_inc(x_356); +lean_dec(x_346); +x_358 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__82; +lean_inc(x_329); lean_inc(x_327); -lean_inc(x_325); -x_357 = l_Lean_mkApp3(x_356, x_325, x_327, x_342); -x_358 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__85; -x_359 = l_Lean_mkApp3(x_358, x_325, x_327, x_354); -x_360 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_361 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_360, x_357); -x_362 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_361, x_359); -x_363 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_363, 0, x_362); -lean_ctor_set(x_363, 1, x_355); -return x_363; +x_359 = l_Lean_mkApp3(x_358, x_327, x_329, x_344); +x_360 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__85; +x_361 = l_Lean_mkApp3(x_360, x_327, x_329, x_356); +x_362 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_363 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_362, x_359); +x_364 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_363, x_361); +x_365 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_365, 0, x_364); +lean_ctor_set(x_365, 1, x_357); +return x_365; } } else { -uint8_t x_364; -lean_dec(x_342); +uint8_t x_366; +lean_dec(x_344); +lean_dec(x_329); lean_dec(x_327); -lean_dec(x_325); -x_364 = !lean_is_exclusive(x_344); -if (x_364 == 0) +x_366 = !lean_is_exclusive(x_346); +if (x_366 == 0) { -return x_344; +return x_346; } else { -lean_object* x_365; lean_object* x_366; lean_object* x_367; -x_365 = lean_ctor_get(x_344, 0); -x_366 = lean_ctor_get(x_344, 1); -lean_inc(x_366); -lean_inc(x_365); -lean_dec(x_344); -x_367 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_367, 0, x_365); -lean_ctor_set(x_367, 1, x_366); -return x_367; +lean_object* x_367; lean_object* x_368; lean_object* x_369; +x_367 = lean_ctor_get(x_346, 0); +x_368 = lean_ctor_get(x_346, 1); +lean_inc(x_368); +lean_inc(x_367); +lean_dec(x_346); +x_369 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_369, 0, x_367); +lean_ctor_set(x_369, 1, x_368); +return x_369; } } } else { -uint8_t x_368; -lean_dec(x_340); +uint8_t x_370; +lean_dec(x_342); +lean_dec(x_329); lean_dec(x_327); -lean_dec(x_325); +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_368 = !lean_is_exclusive(x_341); -if (x_368 == 0) +x_370 = !lean_is_exclusive(x_343); +if (x_370 == 0) { -return x_341; +return x_343; } else { -lean_object* x_369; lean_object* x_370; lean_object* x_371; -x_369 = lean_ctor_get(x_341, 0); -x_370 = lean_ctor_get(x_341, 1); -lean_inc(x_370); -lean_inc(x_369); -lean_dec(x_341); -x_371 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_371, 0, x_369); -lean_ctor_set(x_371, 1, x_370); -return x_371; +lean_object* x_371; lean_object* x_372; lean_object* x_373; +x_371 = lean_ctor_get(x_343, 0); +x_372 = lean_ctor_get(x_343, 1); +lean_inc(x_372); +lean_inc(x_371); +lean_dec(x_343); +x_373 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_373, 0, x_371); +lean_ctor_set(x_373, 1, x_372); +return x_373; } } } } else { -lean_object* x_376; lean_object* x_377; +lean_object* x_378; lean_object* x_379; +lean_dec(x_329); lean_dec(x_327); -lean_dec(x_325); +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_376 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_377 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_377, 0, x_376); -lean_ctor_set(x_377, 1, x_9); -return x_377; +x_378 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_379 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_379, 0, x_378); +lean_ctor_set(x_379, 1, x_11); +return x_379; } } } @@ -5187,629 +5438,616 @@ return x_377; } else { -lean_object* x_378; uint8_t x_379; -lean_dec(x_46); +lean_object* x_380; uint8_t x_381; +lean_dec(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_378 = l_Lean_Elab_Tactic_Omega_natCast_x3f___closed__2; -x_379 = lean_string_dec_eq(x_45, x_378); -lean_dec(x_45); -if (x_379 == 0) +x_380 = l_Lean_Elab_Tactic_Omega_natCast_x3f___closed__2; +x_381 = lean_string_dec_eq(x_47, x_380); +lean_dec(x_47); +if (x_381 == 0) { -lean_object* x_380; lean_object* x_381; -lean_dec(x_44); -x_380 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_381 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_381, 0, x_380); -lean_ctor_set(x_381, 1, x_9); -return x_381; +lean_object* x_382; lean_object* x_383; +lean_dec(x_46); +x_382 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_383 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_383, 0, x_382); +lean_ctor_set(x_383, 1, x_11); +return x_383; } else { -lean_object* x_382; lean_object* x_383; uint8_t x_384; -x_382 = lean_array_get_size(x_44); -x_383 = lean_unsigned_to_nat(3u); -x_384 = lean_nat_dec_eq(x_382, x_383); -lean_dec(x_382); -if (x_384 == 0) +lean_object* x_384; lean_object* x_385; uint8_t x_386; +x_384 = lean_array_get_size(x_46); +x_385 = lean_unsigned_to_nat(3u); +x_386 = lean_nat_dec_eq(x_384, x_385); +lean_dec(x_384); +if (x_386 == 0) { -lean_object* x_385; lean_object* x_386; -lean_dec(x_44); -x_385 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_386 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_386, 0, x_385); -lean_ctor_set(x_386, 1, x_9); -return x_386; +lean_object* x_387; lean_object* x_388; +lean_dec(x_46); +x_387 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_388 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_388, 0, x_387); +lean_ctor_set(x_388, 1, x_11); +return x_388; } else { -lean_object* x_387; lean_object* x_388; -x_387 = lean_unsigned_to_nat(0u); -x_388 = lean_array_fget(x_44, x_387); -if (lean_obj_tag(x_388) == 4) -{ -lean_object* x_389; -x_389 = lean_ctor_get(x_388, 0); -lean_inc(x_389); -if (lean_obj_tag(x_389) == 1) -{ -lean_object* x_390; -x_390 = lean_ctor_get(x_389, 0); -lean_inc(x_390); -if (lean_obj_tag(x_390) == 0) -{ -lean_object* x_391; lean_object* x_392; lean_object* x_393; uint8_t x_394; -x_391 = lean_ctor_get(x_388, 1); +lean_object* x_389; lean_object* x_390; +x_389 = lean_unsigned_to_nat(0u); +x_390 = lean_array_fget(x_46, x_389); +if (lean_obj_tag(x_390) == 4) +{ +lean_object* x_391; +x_391 = lean_ctor_get(x_390, 0); lean_inc(x_391); -lean_dec(x_388); -x_392 = lean_ctor_get(x_389, 1); +if (lean_obj_tag(x_391) == 1) +{ +lean_object* x_392; +x_392 = lean_ctor_get(x_391, 0); lean_inc(x_392); -lean_dec(x_389); -x_393 = l_Lean_Elab_Tactic_Omega_atomsList___rarg___closed__1; -x_394 = lean_string_dec_eq(x_392, x_393); -lean_dec(x_392); -if (x_394 == 0) +if (lean_obj_tag(x_392) == 0) { -lean_object* x_395; lean_object* x_396; +lean_object* x_393; lean_object* x_394; lean_object* x_395; uint8_t x_396; +x_393 = lean_ctor_get(x_390, 1); +lean_inc(x_393); +lean_dec(x_390); +x_394 = lean_ctor_get(x_391, 1); +lean_inc(x_394); lean_dec(x_391); -lean_dec(x_44); -x_395 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_396 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_396, 0, x_395); -lean_ctor_set(x_396, 1, x_9); -return x_396; +x_395 = l_Lean_Elab_Tactic_Omega_atomsList___rarg___closed__1; +x_396 = lean_string_dec_eq(x_394, x_395); +lean_dec(x_394); +if (x_396 == 0) +{ +lean_object* x_397; lean_object* x_398; +lean_dec(x_393); +lean_dec(x_46); +x_397 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_398 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_398, 0, x_397); +lean_ctor_set(x_398, 1, x_11); +return x_398; } else { -if (lean_obj_tag(x_391) == 0) +if (lean_obj_tag(x_393) == 0) { -lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; uint8_t x_403; lean_object* x_404; -x_397 = lean_unsigned_to_nat(2u); -x_398 = lean_array_fget(x_44, x_397); -lean_dec(x_44); -x_399 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__90; -lean_inc(x_398); -x_400 = l_Lean_Expr_app___override(x_399, x_398); -x_401 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; -x_402 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_401, x_400); -x_403 = lean_ctor_get_uint8(x_4, 1); -x_404 = l_Lean_Expr_getAppFnArgs(x_398); -if (x_403 == 0) -{ -lean_object* x_405; -x_405 = lean_ctor_get(x_404, 0); -lean_inc(x_405); -if (lean_obj_tag(x_405) == 1) -{ -lean_object* x_406; -x_406 = lean_ctor_get(x_405, 0); -lean_inc(x_406); -if (lean_obj_tag(x_406) == 1) +lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; uint8_t x_405; lean_object* x_406; +x_399 = lean_unsigned_to_nat(2u); +x_400 = lean_array_fget(x_46, x_399); +lean_dec(x_46); +x_401 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__90; +lean_inc(x_400); +x_402 = l_Lean_Expr_app___override(x_401, x_400); +x_403 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_404 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_403, x_402); +x_405 = lean_ctor_get_uint8(x_4, 1); +x_406 = l_Lean_Expr_getAppFnArgs(x_400); +if (x_405 == 0) { lean_object* x_407; x_407 = lean_ctor_get(x_406, 0); lean_inc(x_407); -if (lean_obj_tag(x_407) == 0) +if (lean_obj_tag(x_407) == 1) { -lean_object* x_408; lean_object* x_409; lean_object* x_410; uint8_t x_411; -x_408 = lean_ctor_get(x_404, 1); +lean_object* x_408; +x_408 = lean_ctor_get(x_407, 0); lean_inc(x_408); -lean_dec(x_404); -x_409 = lean_ctor_get(x_405, 1); +if (lean_obj_tag(x_408) == 1) +{ +lean_object* x_409; +x_409 = lean_ctor_get(x_408, 0); lean_inc(x_409); -lean_dec(x_405); +if (lean_obj_tag(x_409) == 0) +{ +lean_object* x_410; lean_object* x_411; lean_object* x_412; uint8_t x_413; x_410 = lean_ctor_get(x_406, 1); lean_inc(x_410); lean_dec(x_406); -x_411 = lean_string_dec_eq(x_410, x_393); -if (x_411 == 0) -{ -lean_object* x_412; uint8_t x_413; -x_412 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__91; -x_413 = lean_string_dec_eq(x_410, x_412); +x_411 = lean_ctor_get(x_407, 1); +lean_inc(x_411); +lean_dec(x_407); +x_412 = lean_ctor_get(x_408, 1); +lean_inc(x_412); +lean_dec(x_408); +x_413 = lean_string_dec_eq(x_412, x_395); if (x_413 == 0) { lean_object* x_414; uint8_t x_415; -x_414 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__92; -x_415 = lean_string_dec_eq(x_410, x_414); -lean_dec(x_410); +x_414 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__91; +x_415 = lean_string_dec_eq(x_412, x_414); if (x_415 == 0) { -lean_object* x_416; -lean_dec(x_409); -lean_dec(x_408); -x_416 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_416, 0, x_402); -lean_ctor_set(x_416, 1, x_9); -return x_416; +lean_object* x_416; uint8_t x_417; +x_416 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__92; +x_417 = lean_string_dec_eq(x_412, x_416); +lean_dec(x_412); +if (x_417 == 0) +{ +lean_object* x_418; +lean_dec(x_411); +lean_dec(x_410); +x_418 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_418, 0, x_404); +lean_ctor_set(x_418, 1, x_11); +return x_418; } else { -lean_object* x_417; uint8_t x_418; -x_417 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__93; -x_418 = lean_string_dec_eq(x_409, x_417); -lean_dec(x_409); -if (x_418 == 0) +lean_object* x_419; uint8_t x_420; +x_419 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__93; +x_420 = lean_string_dec_eq(x_411, x_419); +lean_dec(x_411); +if (x_420 == 0) { -lean_object* x_419; -lean_dec(x_408); -x_419 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_419, 0, x_402); -lean_ctor_set(x_419, 1, x_9); -return x_419; +lean_object* x_421; +lean_dec(x_410); +x_421 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_421, 0, x_404); +lean_ctor_set(x_421, 1, x_11); +return x_421; } else { -lean_object* x_420; uint8_t x_421; -x_420 = lean_array_get_size(x_408); -x_421 = lean_nat_dec_eq(x_420, x_397); -lean_dec(x_420); -if (x_421 == 0) +lean_object* x_422; uint8_t x_423; +x_422 = lean_array_get_size(x_410); +x_423 = lean_nat_dec_eq(x_422, x_399); +lean_dec(x_422); +if (x_423 == 0) { -lean_object* x_422; -lean_dec(x_408); -x_422 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_422, 0, x_402); -lean_ctor_set(x_422, 1, x_9); -return x_422; +lean_object* x_424; +lean_dec(x_410); +x_424 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_424, 0, x_404); +lean_ctor_set(x_424, 1, x_11); +return x_424; } else { -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; -x_423 = lean_array_fget(x_408, x_387); -x_424 = lean_unsigned_to_nat(1u); -x_425 = lean_array_fget(x_408, x_424); -lean_dec(x_408); -x_426 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__96; -x_427 = l_Lean_mkAppB(x_426, x_423, x_425); -x_428 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_402, x_427); -x_429 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_429, 0, x_428); -lean_ctor_set(x_429, 1, x_9); -return x_429; +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; +x_425 = lean_array_fget(x_410, x_389); +x_426 = lean_unsigned_to_nat(1u); +x_427 = lean_array_fget(x_410, x_426); +lean_dec(x_410); +x_428 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__96; +x_429 = l_Lean_mkAppB(x_428, x_425, x_427); +x_430 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_404, x_429); +x_431 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_431, 0, x_430); +lean_ctor_set(x_431, 1, x_11); +return x_431; } } } } else { -lean_object* x_430; uint8_t x_431; -lean_dec(x_410); -x_430 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__97; -x_431 = lean_string_dec_eq(x_409, x_430); -lean_dec(x_409); -if (x_431 == 0) +lean_object* x_432; uint8_t x_433; +lean_dec(x_412); +x_432 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__97; +x_433 = lean_string_dec_eq(x_411, x_432); +lean_dec(x_411); +if (x_433 == 0) { -lean_object* x_432; -lean_dec(x_408); -x_432 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_432, 0, x_402); -lean_ctor_set(x_432, 1, x_9); -return x_432; +lean_object* x_434; +lean_dec(x_410); +x_434 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_434, 0, x_404); +lean_ctor_set(x_434, 1, x_11); +return x_434; } else { -lean_object* x_433; uint8_t x_434; -x_433 = lean_array_get_size(x_408); -x_434 = lean_nat_dec_eq(x_433, x_397); -lean_dec(x_433); -if (x_434 == 0) +lean_object* x_435; uint8_t x_436; +x_435 = lean_array_get_size(x_410); +x_436 = lean_nat_dec_eq(x_435, x_399); +lean_dec(x_435); +if (x_436 == 0) { -lean_object* x_435; -lean_dec(x_408); -x_435 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_435, 0, x_402); -lean_ctor_set(x_435, 1, x_9); -return x_435; +lean_object* x_437; +lean_dec(x_410); +x_437 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_437, 0, x_404); +lean_ctor_set(x_437, 1, x_11); +return x_437; } else { -lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; -x_436 = lean_array_fget(x_408, x_387); -x_437 = lean_unsigned_to_nat(1u); -x_438 = lean_array_fget(x_408, x_437); -lean_dec(x_408); -x_439 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__99; -x_440 = l_Lean_mkAppB(x_439, x_436, x_438); -x_441 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_402, x_440); -x_442 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_442, 0, x_441); -lean_ctor_set(x_442, 1, x_9); -return x_442; +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; +x_438 = lean_array_fget(x_410, x_389); +x_439 = lean_unsigned_to_nat(1u); +x_440 = lean_array_fget(x_410, x_439); +lean_dec(x_410); +x_441 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__99; +x_442 = l_Lean_mkAppB(x_441, x_438, x_440); +x_443 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_404, x_442); +x_444 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_444, 0, x_443); +lean_ctor_set(x_444, 1, x_11); +return x_444; } } } } else { -lean_object* x_443; uint8_t x_444; -lean_dec(x_410); -x_443 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__100; -x_444 = lean_string_dec_eq(x_409, x_443); -lean_dec(x_409); -if (x_444 == 0) +lean_object* x_445; uint8_t x_446; +lean_dec(x_412); +x_445 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__100; +x_446 = lean_string_dec_eq(x_411, x_445); +lean_dec(x_411); +if (x_446 == 0) { -lean_object* x_445; -lean_dec(x_408); -x_445 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_445, 0, x_402); -lean_ctor_set(x_445, 1, x_9); -return x_445; +lean_object* x_447; +lean_dec(x_410); +x_447 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_447, 0, x_404); +lean_ctor_set(x_447, 1, x_11); +return x_447; } else { -lean_object* x_446; lean_object* x_447; uint8_t x_448; -x_446 = lean_array_get_size(x_408); -x_447 = lean_unsigned_to_nat(1u); -x_448 = lean_nat_dec_eq(x_446, x_447); -lean_dec(x_446); -if (x_448 == 0) +lean_object* x_448; lean_object* x_449; uint8_t x_450; +x_448 = lean_array_get_size(x_410); +x_449 = lean_unsigned_to_nat(1u); +x_450 = lean_nat_dec_eq(x_448, x_449); +lean_dec(x_448); +if (x_450 == 0) { -lean_object* x_449; -lean_dec(x_408); -x_449 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_449, 0, x_402); -lean_ctor_set(x_449, 1, x_9); -return x_449; +lean_object* x_451; +lean_dec(x_410); +x_451 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_451, 0, x_404); +lean_ctor_set(x_451, 1, x_11); +return x_451; } else { -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; -x_450 = lean_array_fget(x_408, x_387); -lean_dec(x_408); -x_451 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__103; -lean_inc(x_450); -x_452 = l_Lean_Expr_app___override(x_451, x_450); -x_453 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__106; -x_454 = l_Lean_Expr_app___override(x_453, x_450); -x_455 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_402, x_452); -x_456 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_455, x_454); -x_457 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_457, 0, x_456); -lean_ctor_set(x_457, 1, x_9); -return x_457; +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; +x_452 = lean_array_fget(x_410, x_389); +lean_dec(x_410); +x_453 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__103; +lean_inc(x_452); +x_454 = l_Lean_Expr_app___override(x_453, x_452); +x_455 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__106; +x_456 = l_Lean_Expr_app___override(x_455, x_452); +x_457 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_404, x_454); +x_458 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_457, x_456); +x_459 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_459, 0, x_458); +lean_ctor_set(x_459, 1, x_11); +return x_459; } } } } else { -lean_object* x_458; +lean_object* x_460; +lean_dec(x_409); +lean_dec(x_408); lean_dec(x_407); lean_dec(x_406); -lean_dec(x_405); -lean_dec(x_404); -x_458 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_458, 0, x_402); -lean_ctor_set(x_458, 1, x_9); -return x_458; +x_460 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_460, 0, x_404); +lean_ctor_set(x_460, 1, x_11); +return x_460; } } else { -lean_object* x_459; +lean_object* x_461; +lean_dec(x_408); +lean_dec(x_407); lean_dec(x_406); -lean_dec(x_405); -lean_dec(x_404); -x_459 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_459, 0, x_402); -lean_ctor_set(x_459, 1, x_9); -return x_459; +x_461 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_461, 0, x_404); +lean_ctor_set(x_461, 1, x_11); +return x_461; } } else { -lean_object* x_460; -lean_dec(x_405); -lean_dec(x_404); -x_460 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_460, 0, x_402); -lean_ctor_set(x_460, 1, x_9); -return x_460; +lean_object* x_462; +lean_dec(x_407); +lean_dec(x_406); +x_462 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_462, 0, x_404); +lean_ctor_set(x_462, 1, x_11); +return x_462; } } else { -lean_object* x_461; -x_461 = lean_ctor_get(x_404, 0); -lean_inc(x_461); -if (lean_obj_tag(x_461) == 1) -{ -lean_object* x_462; -x_462 = lean_ctor_get(x_461, 0); -lean_inc(x_462); -if (lean_obj_tag(x_462) == 1) -{ lean_object* x_463; -x_463 = lean_ctor_get(x_462, 0); +x_463 = lean_ctor_get(x_406, 0); lean_inc(x_463); -if (lean_obj_tag(x_463) == 0) +if (lean_obj_tag(x_463) == 1) { -lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; uint8_t x_468; -x_464 = lean_ctor_get(x_404, 1); +lean_object* x_464; +x_464 = lean_ctor_get(x_463, 0); lean_inc(x_464); -lean_dec(x_404); -x_465 = lean_ctor_get(x_461, 1); +if (lean_obj_tag(x_464) == 1) +{ +lean_object* x_465; +x_465 = lean_ctor_get(x_464, 0); lean_inc(x_465); -lean_dec(x_461); -x_466 = lean_ctor_get(x_462, 1); +if (lean_obj_tag(x_465) == 0) +{ +lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; uint8_t x_470; +x_466 = lean_ctor_get(x_406, 1); lean_inc(x_466); -lean_dec(x_462); -x_467 = l_Lean_Elab_Tactic_Omega_groundNat_x3f___closed__3; -x_468 = lean_string_dec_eq(x_466, x_467); -if (x_468 == 0) -{ -uint8_t x_469; -x_469 = lean_string_dec_eq(x_466, x_393); -if (x_469 == 0) -{ -lean_object* x_470; uint8_t x_471; -x_470 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__91; -x_471 = lean_string_dec_eq(x_466, x_470); +lean_dec(x_406); +x_467 = lean_ctor_get(x_463, 1); +lean_inc(x_467); +lean_dec(x_463); +x_468 = lean_ctor_get(x_464, 1); +lean_inc(x_468); +lean_dec(x_464); +x_469 = l_Lean_Elab_Tactic_Omega_groundNat_x3f___closed__3; +x_470 = lean_string_dec_eq(x_468, x_469); +if (x_470 == 0) +{ +uint8_t x_471; +x_471 = lean_string_dec_eq(x_468, x_395); if (x_471 == 0) { lean_object* x_472; uint8_t x_473; -x_472 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__92; -x_473 = lean_string_dec_eq(x_466, x_472); -lean_dec(x_466); +x_472 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__91; +x_473 = lean_string_dec_eq(x_468, x_472); if (x_473 == 0) { -lean_object* x_474; -lean_dec(x_465); -lean_dec(x_464); -x_474 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_474, 0, x_402); -lean_ctor_set(x_474, 1, x_9); -return x_474; +lean_object* x_474; uint8_t x_475; +x_474 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__92; +x_475 = lean_string_dec_eq(x_468, x_474); +lean_dec(x_468); +if (x_475 == 0) +{ +lean_object* x_476; +lean_dec(x_467); +lean_dec(x_466); +x_476 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_476, 0, x_404); +lean_ctor_set(x_476, 1, x_11); +return x_476; } else { -lean_object* x_475; uint8_t x_476; -x_475 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__93; -x_476 = lean_string_dec_eq(x_465, x_475); -lean_dec(x_465); -if (x_476 == 0) +lean_object* x_477; uint8_t x_478; +x_477 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__93; +x_478 = lean_string_dec_eq(x_467, x_477); +lean_dec(x_467); +if (x_478 == 0) { -lean_object* x_477; -lean_dec(x_464); -x_477 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_477, 0, x_402); -lean_ctor_set(x_477, 1, x_9); -return x_477; +lean_object* x_479; +lean_dec(x_466); +x_479 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_479, 0, x_404); +lean_ctor_set(x_479, 1, x_11); +return x_479; } else { -lean_object* x_478; uint8_t x_479; -x_478 = lean_array_get_size(x_464); -x_479 = lean_nat_dec_eq(x_478, x_397); -lean_dec(x_478); -if (x_479 == 0) +lean_object* x_480; uint8_t x_481; +x_480 = lean_array_get_size(x_466); +x_481 = lean_nat_dec_eq(x_480, x_399); +lean_dec(x_480); +if (x_481 == 0) { -lean_object* x_480; -lean_dec(x_464); -x_480 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_480, 0, x_402); -lean_ctor_set(x_480, 1, x_9); -return x_480; +lean_object* x_482; +lean_dec(x_466); +x_482 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_482, 0, x_404); +lean_ctor_set(x_482, 1, x_11); +return x_482; } else { -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; -x_481 = lean_array_fget(x_464, x_387); -x_482 = lean_unsigned_to_nat(1u); -x_483 = lean_array_fget(x_464, x_482); -lean_dec(x_464); -x_484 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__96; -x_485 = l_Lean_mkAppB(x_484, x_481, x_483); -x_486 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_402, 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_9); -return x_487; +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; +x_483 = lean_array_fget(x_466, x_389); +x_484 = lean_unsigned_to_nat(1u); +x_485 = lean_array_fget(x_466, x_484); +lean_dec(x_466); +x_486 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__96; +x_487 = l_Lean_mkAppB(x_486, x_483, x_485); +x_488 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_404, x_487); +x_489 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_489, 0, x_488); +lean_ctor_set(x_489, 1, x_11); +return x_489; } } } } else { -lean_object* x_488; uint8_t x_489; -lean_dec(x_466); -x_488 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__97; -x_489 = lean_string_dec_eq(x_465, x_488); -lean_dec(x_465); -if (x_489 == 0) +lean_object* x_490; uint8_t x_491; +lean_dec(x_468); +x_490 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__97; +x_491 = lean_string_dec_eq(x_467, x_490); +lean_dec(x_467); +if (x_491 == 0) { -lean_object* x_490; -lean_dec(x_464); -x_490 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_490, 0, x_402); -lean_ctor_set(x_490, 1, x_9); -return x_490; +lean_object* x_492; +lean_dec(x_466); +x_492 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_492, 0, x_404); +lean_ctor_set(x_492, 1, x_11); +return x_492; } else { -lean_object* x_491; uint8_t x_492; -x_491 = lean_array_get_size(x_464); -x_492 = lean_nat_dec_eq(x_491, x_397); -lean_dec(x_491); -if (x_492 == 0) +lean_object* x_493; uint8_t x_494; +x_493 = lean_array_get_size(x_466); +x_494 = lean_nat_dec_eq(x_493, x_399); +lean_dec(x_493); +if (x_494 == 0) { -lean_object* x_493; -lean_dec(x_464); -x_493 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_493, 0, x_402); -lean_ctor_set(x_493, 1, x_9); -return x_493; +lean_object* x_495; +lean_dec(x_466); +x_495 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_495, 0, x_404); +lean_ctor_set(x_495, 1, x_11); +return x_495; } else { -lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; -x_494 = lean_array_fget(x_464, x_387); -x_495 = lean_unsigned_to_nat(1u); -x_496 = lean_array_fget(x_464, x_495); -lean_dec(x_464); -x_497 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__99; -x_498 = l_Lean_mkAppB(x_497, x_494, x_496); -x_499 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_402, x_498); -x_500 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_500, 0, x_499); -lean_ctor_set(x_500, 1, x_9); -return x_500; +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; +x_496 = lean_array_fget(x_466, x_389); +x_497 = lean_unsigned_to_nat(1u); +x_498 = lean_array_fget(x_466, x_497); +lean_dec(x_466); +x_499 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__99; +x_500 = l_Lean_mkAppB(x_499, x_496, x_498); +x_501 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_404, 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_11); +return x_502; } } } } else { -lean_object* x_501; uint8_t x_502; -lean_dec(x_466); -x_501 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__100; -x_502 = lean_string_dec_eq(x_465, x_501); -lean_dec(x_465); -if (x_502 == 0) +lean_object* x_503; uint8_t x_504; +lean_dec(x_468); +x_503 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__100; +x_504 = lean_string_dec_eq(x_467, x_503); +lean_dec(x_467); +if (x_504 == 0) { -lean_object* x_503; -lean_dec(x_464); -x_503 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_503, 0, x_402); -lean_ctor_set(x_503, 1, x_9); -return x_503; +lean_object* x_505; +lean_dec(x_466); +x_505 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_505, 0, x_404); +lean_ctor_set(x_505, 1, x_11); +return x_505; } else { -lean_object* x_504; lean_object* x_505; uint8_t x_506; -x_504 = lean_array_get_size(x_464); -x_505 = lean_unsigned_to_nat(1u); -x_506 = lean_nat_dec_eq(x_504, x_505); -lean_dec(x_504); -if (x_506 == 0) +lean_object* x_506; lean_object* x_507; uint8_t x_508; +x_506 = lean_array_get_size(x_466); +x_507 = lean_unsigned_to_nat(1u); +x_508 = lean_nat_dec_eq(x_506, x_507); +lean_dec(x_506); +if (x_508 == 0) { -lean_object* x_507; -lean_dec(x_464); -x_507 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_507, 0, x_402); -lean_ctor_set(x_507, 1, x_9); -return x_507; +lean_object* x_509; +lean_dec(x_466); +x_509 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_509, 0, x_404); +lean_ctor_set(x_509, 1, x_11); +return x_509; } else { -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; -x_508 = lean_array_fget(x_464, x_387); -lean_dec(x_464); -x_509 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__103; -lean_inc(x_508); -x_510 = l_Lean_Expr_app___override(x_509, x_508); -x_511 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__106; -x_512 = l_Lean_Expr_app___override(x_511, x_508); -x_513 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_402, x_510); -x_514 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_513, x_512); -x_515 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_515, 0, x_514); -lean_ctor_set(x_515, 1, x_9); -return x_515; +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; +x_510 = lean_array_fget(x_466, x_389); +lean_dec(x_466); +x_511 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__103; +lean_inc(x_510); +x_512 = l_Lean_Expr_app___override(x_511, x_510); +x_513 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__106; +x_514 = l_Lean_Expr_app___override(x_513, x_510); +x_515 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_404, x_512); +x_516 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_515, x_514); +x_517 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_517, 0, x_516); +lean_ctor_set(x_517, 1, x_11); +return x_517; } } } } else { -lean_object* x_516; uint8_t x_517; -lean_dec(x_466); -x_516 = l_Lean_Elab_Tactic_Omega_groundNat_x3f___closed__10; -x_517 = lean_string_dec_eq(x_465, x_516); -lean_dec(x_465); -if (x_517 == 0) +lean_object* x_518; uint8_t x_519; +lean_dec(x_468); +x_518 = l_Lean_Elab_Tactic_Omega_groundNat_x3f___closed__10; +x_519 = lean_string_dec_eq(x_467, x_518); +lean_dec(x_467); +if (x_519 == 0) { -lean_object* x_518; -lean_dec(x_464); -x_518 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_518, 0, x_402); -lean_ctor_set(x_518, 1, x_9); -return x_518; +lean_object* x_520; +lean_dec(x_466); +x_520 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_520, 0, x_404); +lean_ctor_set(x_520, 1, x_11); +return x_520; } else { -lean_object* x_519; lean_object* x_520; uint8_t x_521; -x_519 = lean_array_get_size(x_464); -x_520 = lean_unsigned_to_nat(6u); -x_521 = lean_nat_dec_eq(x_519, x_520); -lean_dec(x_519); -if (x_521 == 0) +lean_object* x_521; lean_object* x_522; uint8_t x_523; +x_521 = lean_array_get_size(x_466); +x_522 = lean_unsigned_to_nat(6u); +x_523 = lean_nat_dec_eq(x_521, x_522); +lean_dec(x_521); +if (x_523 == 0) { -lean_object* x_522; -lean_dec(x_464); -x_522 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_522, 0, x_402); -lean_ctor_set(x_522, 1, x_9); -return x_522; +lean_object* x_524; +lean_dec(x_466); +x_524 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_524, 0, x_404); +lean_ctor_set(x_524, 1, x_11); +return x_524; } else { -lean_object* x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; -x_523 = lean_unsigned_to_nat(4u); -x_524 = lean_array_fget(x_464, x_523); -x_525 = lean_unsigned_to_nat(5u); -x_526 = lean_array_fget(x_464, x_525); -lean_dec(x_464); -x_527 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__109; -x_528 = l_Lean_mkAppB(x_527, x_524, x_526); -x_529 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_402, 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_9); -return x_530; +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_525 = lean_unsigned_to_nat(4u); +x_526 = lean_array_fget(x_466, x_525); +x_527 = lean_unsigned_to_nat(5u); +x_528 = lean_array_fget(x_466, x_527); +lean_dec(x_466); +x_529 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__109; +x_530 = l_Lean_mkAppB(x_529, x_526, x_528); +x_531 = l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(x_404, 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_11); +return x_532; } } } } else { -lean_object* x_531; +lean_object* x_533; +lean_dec(x_465); +lean_dec(x_464); lean_dec(x_463); -lean_dec(x_462); -lean_dec(x_461); -lean_dec(x_404); -x_531 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_531, 0, x_402); -lean_ctor_set(x_531, 1, x_9); -return x_531; -} -} -else -{ -lean_object* x_532; -lean_dec(x_462); -lean_dec(x_461); -lean_dec(x_404); -x_532 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_532, 0, x_402); -lean_ctor_set(x_532, 1, x_9); -return x_532; +lean_dec(x_406); +x_533 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_533, 0, x_404); +lean_ctor_set(x_533, 1, x_11); +return x_533; } } else { -lean_object* x_533; -lean_dec(x_461); -lean_dec(x_404); -x_533 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_533, 0, x_402); -lean_ctor_set(x_533, 1, x_9); -return x_533; -} +lean_object* x_534; +lean_dec(x_464); +lean_dec(x_463); +lean_dec(x_406); +x_534 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_534, 0, x_404); +lean_ctor_set(x_534, 1, x_11); +return x_534; } } else { -lean_object* x_534; lean_object* x_535; -lean_dec(x_391); -lean_dec(x_44); -x_534 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +lean_object* x_535; +lean_dec(x_463); +lean_dec(x_406); x_535 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_535, 0, x_534); -lean_ctor_set(x_535, 1, x_9); +lean_ctor_set(x_535, 0, x_404); +lean_ctor_set(x_535, 1, x_11); return x_535; } } @@ -5817,107 +6055,123 @@ return x_535; else { lean_object* x_536; lean_object* x_537; -lean_dec(x_390); -lean_dec(x_389); -lean_dec(x_388); -lean_dec(x_44); +lean_dec(x_393); +lean_dec(x_46); x_536 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; x_537 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_537, 0, x_536); -lean_ctor_set(x_537, 1, x_9); +lean_ctor_set(x_537, 1, x_11); return x_537; } } +} else { lean_object* x_538; lean_object* x_539; -lean_dec(x_389); -lean_dec(x_388); -lean_dec(x_44); +lean_dec(x_392); +lean_dec(x_391); +lean_dec(x_390); +lean_dec(x_46); x_538 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; x_539 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_539, 0, x_538); -lean_ctor_set(x_539, 1, x_9); +lean_ctor_set(x_539, 1, x_11); return x_539; } } else { lean_object* x_540; lean_object* x_541; -lean_dec(x_388); -lean_dec(x_44); +lean_dec(x_391); +lean_dec(x_390); +lean_dec(x_46); x_540 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; x_541 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_541, 0, x_540); -lean_ctor_set(x_541, 1, x_9); +lean_ctor_set(x_541, 1, x_11); return x_541; } } -} -} -} else { lean_object* x_542; lean_object* x_543; -lean_dec(x_43); -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_390); +lean_dec(x_46); x_542 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; x_543 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_543, 0, x_542); -lean_ctor_set(x_543, 1, x_9); +lean_ctor_set(x_543, 1, x_11); return x_543; } } -default: +} +} +} +else { lean_object* x_544; lean_object* x_545; +lean_dec(x_45); +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); x_544 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; x_545 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_545, 0, x_544); -lean_ctor_set(x_545, 1, x_9); +lean_ctor_set(x_545, 1, x_11); return x_545; } } -} -else +default: { lean_object* x_546; lean_object* x_547; -lean_dec(x_11); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_10); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); x_546 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; x_547 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_547, 0, x_546); -lean_ctor_set(x_547, 1, x_9); +lean_ctor_set(x_547, 1, x_11); return x_547; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +else +{ +lean_object* x_548; lean_object* x_549; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +x_548 = l_Lean_Elab_Tactic_Omega_analyzeAtom___closed__1; +x_549 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_549, 0, x_548); +lean_ctor_set(x_549, 1, x_11); +return x_549; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_analyzeAtom___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_Elab_Tactic_Omega_analyzeAtom(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +uint8_t x_12; lean_object* x_13; +x_12 = lean_unbox(x_5); +lean_dec(x_5); +x_13 = l_Lean_Elab_Tactic_Omega_analyzeAtom(x_1, x_2, x_3, x_4, x_12, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_10; +return x_13; } } LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_Elab_Tactic_Omega_lookup___spec__2(lean_object* x_1, lean_object* x_2) { @@ -5978,40 +6232,40 @@ x_1 = l_Lean_inheritedTraceOptions; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__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) { _start: { -lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3___closed__1; -x_11 = lean_st_ref_get(x_10, x_9); -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) +lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_12 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3___closed__1; +x_13 = lean_st_ref_get(x_12, x_11); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) { -lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; -x_13 = lean_ctor_get(x_11, 0); -x_14 = lean_ctor_get(x_7, 2); -x_15 = l___private_Lean_Util_Trace_0__Lean_checkTraceOption(x_13, x_14, x_1); -lean_dec(x_13); -x_16 = lean_box(x_15); -lean_ctor_set(x_11, 0, x_16); -return x_11; +lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_ctor_get(x_9, 2); +x_17 = l___private_Lean_Util_Trace_0__Lean_checkTraceOption(x_15, x_16, x_1); +lean_dec(x_15); +x_18 = lean_box(x_17); +lean_ctor_set(x_13, 0, x_18); +return x_13; } 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; -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_ctor_get(x_7, 2); -x_20 = l___private_Lean_Util_Trace_0__Lean_checkTraceOption(x_17, x_19, x_1); -lean_dec(x_17); -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_18); -return x_22; +lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_23; lean_object* x_24; +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_ctor_get(x_9, 2); +x_22 = l___private_Lean_Util_Trace_0__Lean_checkTraceOption(x_19, x_21, x_1); +lean_dec(x_19); +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_20); +return x_24; } } } @@ -6431,313 +6685,313 @@ return x_10; } } } -LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Elab_Tactic_Omega_lookup___spec__13(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Elab_Tactic_Omega_lookup___spec__13(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_object* x_11, lean_object* x_12) { _start: { if (lean_obj_tag(x_1) == 0) { -lean_object* x_11; lean_object* x_12; +lean_object* x_13; lean_object* x_14; +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -x_11 = l_List_reverse___rarg(x_2); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_10); -return x_12; +x_13 = l_List_reverse___rarg(x_2); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_12); +return x_14; } else { -uint8_t x_13; -x_13 = !lean_is_exclusive(x_1); -if (x_13 == 0) +uint8_t x_15; +x_15 = !lean_is_exclusive(x_1); +if (x_15 == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_1, 0); -x_15 = lean_ctor_get(x_1, 1); +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_1, 0); +x_17 = lean_ctor_get(x_1, 1); +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_16 = lean_infer_type(x_14, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_16) == 0) +x_18 = lean_infer_type(x_16, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_17; lean_object* x_18; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); -lean_inc(x_18); -lean_dec(x_16); +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); lean_ctor_set(x_1, 1, x_2); -lean_ctor_set(x_1, 0, x_17); +lean_ctor_set(x_1, 0, x_19); { -lean_object* _tmp_0 = x_15; +lean_object* _tmp_0 = x_17; lean_object* _tmp_1 = x_1; -lean_object* _tmp_9 = x_18; +lean_object* _tmp_11 = x_20; x_1 = _tmp_0; x_2 = _tmp_1; -x_10 = _tmp_9; +x_12 = _tmp_11; } goto _start; } else { -uint8_t x_20; +uint8_t x_22; lean_free_object(x_1); -lean_dec(x_15); +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); -x_20 = !lean_is_exclusive(x_16); -if (x_20 == 0) +x_22 = !lean_is_exclusive(x_18); +if (x_22 == 0) { -return x_16; +return x_18; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_16, 0); -x_22 = lean_ctor_get(x_16, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_16); -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_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_18, 0); +x_24 = lean_ctor_get(x_18, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_18); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; } } } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_1, 0); -x_25 = lean_ctor_get(x_1, 1); -lean_inc(x_25); -lean_inc(x_24); +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_1, 0); +x_27 = lean_ctor_get(x_1, 1); +lean_inc(x_27); +lean_inc(x_26); lean_dec(x_1); +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_26 = lean_infer_type(x_24, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_26) == 0) +x_28 = lean_infer_type(x_26, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_28) == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_26, 1); -lean_inc(x_28); -lean_dec(x_26); -x_29 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_2); -x_1 = x_25; -x_2 = x_29; -x_10 = x_28; +lean_object* x_29; lean_object* x_30; lean_object* x_31; +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_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_2); +x_1 = x_27; +x_2 = x_31; +x_12 = x_30; goto _start; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; -lean_dec(x_25); +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +lean_dec(x_27); +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_31 = lean_ctor_get(x_26, 0); -lean_inc(x_31); -x_32 = lean_ctor_get(x_26, 1); -lean_inc(x_32); -if (lean_is_exclusive(x_26)) { - lean_ctor_release(x_26, 0); - lean_ctor_release(x_26, 1); - x_33 = x_26; +x_33 = lean_ctor_get(x_28, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_28, 1); +lean_inc(x_34); +if (lean_is_exclusive(x_28)) { + lean_ctor_release(x_28, 0); + lean_ctor_release(x_28, 1); + x_35 = x_28; } else { - lean_dec_ref(x_26); - x_33 = lean_box(0); + lean_dec_ref(x_28); + x_35 = lean_box(0); } -if (lean_is_scalar(x_33)) { - x_34 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_35)) { + x_36 = lean_alloc_ctor(1, 2, 0); } else { - x_34 = x_33; + x_36 = x_35; } -lean_ctor_set(x_34, 0, x_31); -lean_ctor_set(x_34, 1, x_32); -return x_34; +lean_ctor_set(x_36, 0, x_33); +lean_ctor_set(x_36, 1, x_34); +return x_36; } } } } } -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(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_object* x_11, lean_object* x_12) { _start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_11 = lean_ctor_get(x_8, 5); -x_12 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_2, x_6, x_7, x_8, x_9, 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_st_ref_take(x_9, x_14); -x_16 = lean_ctor_get(x_15, 0); +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_13 = lean_ctor_get(x_10, 5); +x_14 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_2, x_8, x_9, x_10, x_11, x_12); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); lean_inc(x_16); -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_dec(x_14); +x_17 = lean_st_ref_take(x_11, 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_is_exclusive(x_18); +if (x_20 == 0) { -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; uint8_t x_26; -x_19 = lean_ctor_get(x_16, 3); -x_20 = l_Lean_HashMap_toArray___at_Lean_Elab_Tactic_Omega_atoms___spec__1___closed__1; -x_21 = 0; -x_22 = lean_alloc_ctor(9, 3, 1); -lean_ctor_set(x_22, 0, x_1); -lean_ctor_set(x_22, 1, x_13); -lean_ctor_set(x_22, 2, x_20); -lean_ctor_set_uint8(x_22, sizeof(void*)*3, x_21); -lean_inc(x_11); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_11); -lean_ctor_set(x_23, 1, x_22); -x_24 = l_Lean_PersistentArray_push___rarg(x_19, x_23); -lean_ctor_set(x_16, 3, x_24); -x_25 = lean_st_ref_set(x_9, x_16, x_17); -x_26 = !lean_is_exclusive(x_25); -if (x_26 == 0) -{ -lean_object* x_27; lean_object* x_28; -x_27 = lean_ctor_get(x_25, 0); -lean_dec(x_27); -x_28 = lean_box(0); -lean_ctor_set(x_25, 0, x_28); -return x_25; +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; uint8_t x_28; +x_21 = lean_ctor_get(x_18, 3); +x_22 = l_Lean_HashMap_toArray___at_Lean_Elab_Tactic_Omega_atoms___spec__1___closed__1; +x_23 = 0; +x_24 = lean_alloc_ctor(9, 3, 1); +lean_ctor_set(x_24, 0, x_1); +lean_ctor_set(x_24, 1, x_15); +lean_ctor_set(x_24, 2, x_22); +lean_ctor_set_uint8(x_24, sizeof(void*)*3, x_23); +lean_inc(x_13); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_13); +lean_ctor_set(x_25, 1, x_24); +x_26 = l_Lean_PersistentArray_push___rarg(x_21, x_25); +lean_ctor_set(x_18, 3, x_26); +x_27 = lean_st_ref_set(x_11, x_18, x_19); +x_28 = !lean_is_exclusive(x_27); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; +x_29 = lean_ctor_get(x_27, 0); +lean_dec(x_29); +x_30 = lean_box(0); +lean_ctor_set(x_27, 0, x_30); +return x_27; } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_25, 1); -lean_inc(x_29); -lean_dec(x_25); -x_30 = lean_box(0); -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; +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_27, 1); +lean_inc(x_31); +lean_dec(x_27); +x_32 = lean_box(0); +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 { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; uint8_t x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_32 = lean_ctor_get(x_16, 0); -x_33 = lean_ctor_get(x_16, 1); -x_34 = lean_ctor_get(x_16, 2); -x_35 = lean_ctor_get(x_16, 3); -x_36 = lean_ctor_get(x_16, 4); -x_37 = lean_ctor_get(x_16, 5); -x_38 = lean_ctor_get(x_16, 6); +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; 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_34 = lean_ctor_get(x_18, 0); +x_35 = lean_ctor_get(x_18, 1); +x_36 = lean_ctor_get(x_18, 2); +x_37 = lean_ctor_get(x_18, 3); +x_38 = lean_ctor_get(x_18, 4); +x_39 = lean_ctor_get(x_18, 5); +x_40 = lean_ctor_get(x_18, 6); +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_inc(x_34); -lean_inc(x_33); -lean_inc(x_32); -lean_dec(x_16); -x_39 = l_Lean_HashMap_toArray___at_Lean_Elab_Tactic_Omega_atoms___spec__1___closed__1; -x_40 = 0; -x_41 = lean_alloc_ctor(9, 3, 1); -lean_ctor_set(x_41, 0, x_1); -lean_ctor_set(x_41, 1, x_13); -lean_ctor_set(x_41, 2, x_39); -lean_ctor_set_uint8(x_41, sizeof(void*)*3, x_40); -lean_inc(x_11); -x_42 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_42, 0, x_11); -lean_ctor_set(x_42, 1, x_41); -x_43 = l_Lean_PersistentArray_push___rarg(x_35, x_42); -x_44 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_44, 0, x_32); -lean_ctor_set(x_44, 1, x_33); -lean_ctor_set(x_44, 2, x_34); -lean_ctor_set(x_44, 3, x_43); -lean_ctor_set(x_44, 4, x_36); -lean_ctor_set(x_44, 5, x_37); -lean_ctor_set(x_44, 6, x_38); -x_45 = lean_st_ref_set(x_9, x_44, x_17); -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; +lean_dec(x_18); +x_41 = l_Lean_HashMap_toArray___at_Lean_Elab_Tactic_Omega_atoms___spec__1___closed__1; +x_42 = 0; +x_43 = lean_alloc_ctor(9, 3, 1); +lean_ctor_set(x_43, 0, x_1); +lean_ctor_set(x_43, 1, x_15); +lean_ctor_set(x_43, 2, x_41); +lean_ctor_set_uint8(x_43, sizeof(void*)*3, x_42); +lean_inc(x_13); +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_13); +lean_ctor_set(x_44, 1, x_43); +x_45 = l_Lean_PersistentArray_push___rarg(x_37, x_44); +x_46 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_46, 0, x_34); +lean_ctor_set(x_46, 1, x_35); +lean_ctor_set(x_46, 2, x_36); +lean_ctor_set(x_46, 3, x_45); +lean_ctor_set(x_46, 4, x_38); +lean_ctor_set(x_46, 5, x_39); +lean_ctor_set(x_46, 6, x_40); +x_47 = lean_st_ref_set(x_11, x_46, x_19); +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_45); - x_47 = lean_box(0); + lean_dec_ref(x_47); + x_49 = lean_box(0); } -x_48 = lean_box(0); -if (lean_is_scalar(x_47)) { - x_49 = lean_alloc_ctor(0, 2, 0); +x_50 = lean_box(0); +if (lean_is_scalar(x_49)) { + x_51 = lean_alloc_ctor(0, 2, 0); } else { - x_49 = x_47; + x_51 = x_49; } -lean_ctor_set(x_49, 0, x_48); -lean_ctor_set(x_49, 1, x_46); -return x_49; +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_48); +return x_51; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_lookup___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_Lean_Elab_Tactic_Omega_lookup___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, 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) { _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; -x_12 = lean_st_ref_take(x_5, x_11); -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_12, 1); -lean_inc(x_14); -lean_dec(x_12); -x_15 = lean_ctor_get(x_13, 0); -lean_inc(x_15); +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_14 = lean_st_ref_take(x_5, x_13); +x_15 = lean_ctor_get(x_14, 0); lean_inc(x_15); -x_16 = l_Lean_HashMap_insert___at_Lean_Elab_Tactic_Omega_lookup___spec__4(x_13, x_1, x_15); -x_17 = lean_st_ref_set(x_5, x_16, x_14); -x_18 = !lean_is_exclusive(x_17); -if (x_18 == 0) +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = lean_ctor_get(x_15, 0); +lean_inc(x_17); +lean_inc(x_17); +x_18 = l_Lean_HashMap_insert___at_Lean_Elab_Tactic_Omega_lookup___spec__4(x_15, x_1, x_17); +x_19 = lean_st_ref_set(x_5, x_18, x_16); +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_17, 0); -lean_dec(x_19); -x_20 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_20, 0, x_2); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_15); -lean_ctor_set(x_21, 1, x_20); -lean_ctor_set(x_17, 0, x_21); -return x_17; +lean_object* x_21; lean_object* x_22; lean_object* x_23; +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_2); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_17); +lean_ctor_set(x_23, 1, x_22); +lean_ctor_set(x_19, 0, x_23); +return x_19; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_22 = lean_ctor_get(x_17, 1); -lean_inc(x_22); -lean_dec(x_17); -x_23 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_23, 0, x_2); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_15); -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_22); -return x_25; +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_24 = lean_ctor_get(x_19, 1); +lean_inc(x_24); +lean_dec(x_19); +x_25 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_25, 0, x_2); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_17); +lean_ctor_set(x_26, 1, 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; } } } @@ -6775,217 +7029,223 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_lookup___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_lookup___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, lean_object* x_13, lean_object* x_14) { _start: { -lean_object* x_13; +lean_object* x_15; lean_dec(x_4); +lean_inc(x_13); +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); lean_inc(x_1); -x_13 = l_Lean_Elab_Tactic_Omega_analyzeAtom(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_13) == 0) +x_15 = l_Lean_Elab_Tactic_Omega_analyzeAtom(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); -x_16 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_15); -x_17 = lean_ctor_get(x_16, 0); +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); lean_inc(x_17); -x_18 = lean_unbox(x_17); -lean_dec(x_17); -if (x_18 == 0) +lean_dec(x_15); +x_18 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_17); +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_19; lean_object* x_20; lean_object* x_21; +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_dec(x_3); -x_19 = lean_ctor_get(x_16, 1); -lean_inc(x_19); -lean_dec(x_16); -x_20 = lean_box(0); -x_21 = l_Lean_Elab_Tactic_Omega_lookup___lambda__1(x_1, x_14, x_20, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_19); +x_21 = lean_ctor_get(x_18, 1); +lean_inc(x_21); +lean_dec(x_18); +x_22 = lean_box(0); +x_23 = l_Lean_Elab_Tactic_Omega_lookup___lambda__1(x_1, x_16, x_22, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_21); +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); -return x_21; +return x_23; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_22 = lean_ctor_get(x_16, 1); -lean_inc(x_22); -lean_dec(x_16); -x_23 = lean_ctor_get(x_14, 0); -lean_inc(x_23); -x_24 = lean_unsigned_to_nat(0u); -x_25 = lean_nat_dec_eq(x_23, x_24); -lean_dec(x_23); -if (x_25 == 0) +lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_24 = lean_ctor_get(x_18, 1); +lean_inc(x_24); +lean_dec(x_18); +x_25 = lean_ctor_get(x_16, 0); +lean_inc(x_25); +x_26 = lean_unsigned_to_nat(0u); +x_27 = lean_nat_dec_eq(x_25, x_26); +lean_dec(x_25); +if (x_27 == 0) { -lean_object* x_26; lean_object* x_27; uint8_t x_28; +lean_object* x_28; lean_object* x_29; uint8_t x_30; lean_inc(x_3); -x_26 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_22); -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_unbox(x_27); -lean_dec(x_27); -if (x_28 == 0) +x_28 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_24); +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_unbox(x_29); +lean_dec(x_29); +if (x_30 == 0) { -lean_object* x_29; lean_object* x_30; lean_object* x_31; +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_dec(x_3); -x_29 = lean_ctor_get(x_26, 1); -lean_inc(x_29); -lean_dec(x_26); -x_30 = lean_box(0); -x_31 = l_Lean_Elab_Tactic_Omega_lookup___lambda__1(x_1, x_14, x_30, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_29); +x_31 = lean_ctor_get(x_28, 1); +lean_inc(x_31); +lean_dec(x_28); +x_32 = lean_box(0); +x_33 = l_Lean_Elab_Tactic_Omega_lookup___lambda__1(x_1, x_16, x_32, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_31); +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); -return x_31; +return x_33; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_32 = lean_ctor_get(x_26, 1); -lean_inc(x_32); -lean_dec(x_26); -lean_inc(x_14); -x_33 = l_Lean_HashSet_toList___at_Lean_Elab_Tactic_Omega_lookup___spec__10(x_14); -x_34 = lean_box(0); +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_34 = lean_ctor_get(x_28, 1); +lean_inc(x_34); +lean_dec(x_28); +lean_inc(x_16); +x_35 = l_Lean_HashSet_toList___at_Lean_Elab_Tactic_Omega_lookup___spec__10(x_16); +x_36 = lean_box(0); +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_35 = l_List_mapM_loop___at_Lean_Elab_Tactic_Omega_lookup___spec__13(x_33, x_34, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_32); -if (lean_obj_tag(x_35) == 0) +x_37 = l_List_mapM_loop___at_Lean_Elab_Tactic_Omega_lookup___spec__13(x_35, x_36, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_34); +if (lean_obj_tag(x_37) == 0) { -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; 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_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_List_mapTR_loop___at_Lean_MessageData_instCoeListExprMessageData___spec__1(x_36, x_34); -x_39 = l_Lean_MessageData_ofList(x_38); -lean_dec(x_38); -x_40 = l_Lean_Elab_Tactic_Omega_lookup___lambda__2___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_Elab_Tactic_Omega_lookup___lambda__2___closed__4; +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +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 = l_List_mapTR_loop___at_Lean_MessageData_instCoeListExprMessageData___spec__1(x_38, x_36); +x_41 = l_Lean_MessageData_ofList(x_40); +lean_dec(x_40); +x_42 = l_Lean_Elab_Tactic_Omega_lookup___lambda__2___closed__2; 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_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_3, x_43, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_37); -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_Elab_Tactic_Omega_lookup___lambda__1(x_1, x_14, x_45, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_46); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_41); +x_44 = l_Lean_Elab_Tactic_Omega_lookup___lambda__2___closed__4; +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 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_3, x_45, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_39); +x_47 = lean_ctor_get(x_46, 0); +lean_inc(x_47); +x_48 = lean_ctor_get(x_46, 1); +lean_inc(x_48); +lean_dec(x_46); +x_49 = l_Lean_Elab_Tactic_Omega_lookup___lambda__1(x_1, x_16, x_47, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_48); +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_45); -return x_47; +lean_dec(x_47); +return x_49; } else { -uint8_t x_48; -lean_dec(x_14); +uint8_t x_50; +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_3); lean_dec(x_1); -x_48 = !lean_is_exclusive(x_35); -if (x_48 == 0) +x_50 = !lean_is_exclusive(x_37); +if (x_50 == 0) { -return x_35; +return x_37; } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_49 = lean_ctor_get(x_35, 0); -x_50 = lean_ctor_get(x_35, 1); -lean_inc(x_50); -lean_inc(x_49); -lean_dec(x_35); -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_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_37, 0); +x_52 = lean_ctor_get(x_37, 1); +lean_inc(x_52); +lean_inc(x_51); +lean_dec(x_37); +x_53 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_52); +return x_53; } } } } else { -lean_object* x_52; lean_object* x_53; +lean_object* x_54; lean_object* x_55; lean_dec(x_3); -x_52 = lean_box(0); -x_53 = l_Lean_Elab_Tactic_Omega_lookup___lambda__1(x_1, x_14, x_52, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_22); +x_54 = lean_box(0); +x_55 = l_Lean_Elab_Tactic_Omega_lookup___lambda__1(x_1, x_16, x_54, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_24); +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); 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_53; +return x_55; } } } else { -uint8_t x_54; +uint8_t x_56; +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_54 = !lean_is_exclusive(x_13); -if (x_54 == 0) +x_56 = !lean_is_exclusive(x_15); +if (x_56 == 0) { -return x_13; +return x_15; } else { -lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_55 = lean_ctor_get(x_13, 0); -x_56 = lean_ctor_get(x_13, 1); -lean_inc(x_56); -lean_inc(x_55); -lean_dec(x_13); -x_57 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_57, 0, x_55); -lean_ctor_set(x_57, 1, x_56); -return x_57; +lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_57 = lean_ctor_get(x_15, 0); +x_58 = lean_ctor_get(x_15, 1); +lean_inc(x_58); +lean_inc(x_57); +lean_dec(x_15); +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; } } } @@ -7025,164 +7285,212 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_lookup(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_lookup(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) { _start: { -lean_object* x_10; uint8_t x_11; -x_10 = lean_st_ref_get(x_3, 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; lean_object* x_15; +x_12 = lean_st_ref_get(x_3, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_15 = l_Lean_Meta_Canonicalizer_canon(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_10, 0); -x_13 = lean_ctor_get(x_10, 1); -lean_inc(x_1); -x_14 = l_Lean_HashMapImp_find_x3f___at_Lean_Elab_Tactic_Omega_lookup___spec__1(x_12, x_1); -if (lean_obj_tag(x_14) == 0) +uint8_t x_16; +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -lean_free_object(x_10); -x_15 = l_Lean_Elab_Tactic_Omega_lookup___closed__2; -x_16 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_13); -x_17 = lean_ctor_get(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_ctor_get(x_15, 1); lean_inc(x_17); -x_18 = lean_unbox(x_17); -lean_dec(x_17); -if (x_18 == 0) +x_19 = l_Lean_HashMapImp_find_x3f___at_Lean_Elab_Tactic_Omega_lookup___spec__1(x_13, x_17); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +lean_free_object(x_15); +x_20 = l_Lean_Elab_Tactic_Omega_lookup___closed__2; +x_21 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_20, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_18); +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) { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_16, 1); -lean_inc(x_19); -lean_dec(x_16); -x_20 = lean_box(0); -x_21 = l_Lean_Elab_Tactic_Omega_lookup___lambda__2(x_1, x_15, x_15, x_20, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_19); -return x_21; +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_21, 1); +lean_inc(x_24); +lean_dec(x_21); +x_25 = lean_box(0); +x_26 = l_Lean_Elab_Tactic_Omega_lookup___lambda__2(x_17, x_20, x_20, x_25, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_24); +return x_26; } 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; -x_22 = lean_ctor_get(x_16, 1); -lean_inc(x_22); -lean_dec(x_16); -lean_inc(x_1); -x_23 = l_Lean_MessageData_ofExpr(x_1); -x_24 = l_Lean_Elab_Tactic_Omega_lookup___closed__4; -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_Tactic_Omega_lookup___lambda__2___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_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_15, x_27, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_22); -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_Lean_Elab_Tactic_Omega_lookup___lambda__2(x_1, x_15, x_15, x_29, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_30); -return x_31; +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_27 = lean_ctor_get(x_21, 1); +lean_inc(x_27); +lean_dec(x_21); +lean_inc(x_17); +x_28 = l_Lean_MessageData_ofExpr(x_17); +x_29 = l_Lean_Elab_Tactic_Omega_lookup___closed__4; +x_30 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_28); +x_31 = l_Lean_Elab_Tactic_Omega_lookup___lambda__2___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 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_20, x_32, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_27); +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_Lean_Elab_Tactic_Omega_lookup___lambda__2(x_17, x_20, x_20, x_34, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_35); +return x_36; } } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; +lean_object* x_37; lean_object* x_38; lean_object* x_39; +lean_dec(x_17); +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_32 = lean_ctor_get(x_14, 0); -lean_inc(x_32); -lean_dec(x_14); -x_33 = lean_box(0); -x_34 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_34, 0, x_32); -lean_ctor_set(x_34, 1, x_33); -lean_ctor_set(x_10, 0, x_34); -return x_10; +x_37 = lean_ctor_get(x_19, 0); +lean_inc(x_37); +lean_dec(x_19); +x_38 = lean_box(0); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +lean_ctor_set(x_15, 0, x_39); +return x_15; } } 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); -lean_inc(x_1); -x_37 = l_Lean_HashMapImp_find_x3f___at_Lean_Elab_Tactic_Omega_lookup___spec__1(x_35, x_1); -if (lean_obj_tag(x_37) == 0) -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; -x_38 = l_Lean_Elab_Tactic_Omega_lookup___closed__2; -x_39 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_38, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_36); -x_40 = lean_ctor_get(x_39, 0); +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_15, 0); +x_41 = lean_ctor_get(x_15, 1); +lean_inc(x_41); lean_inc(x_40); -x_41 = lean_unbox(x_40); -lean_dec(x_40); -if (x_41 == 0) +lean_dec(x_15); +lean_inc(x_40); +x_42 = l_Lean_HashMapImp_find_x3f___at_Lean_Elab_Tactic_Omega_lookup___spec__1(x_13, x_40); +if (lean_obj_tag(x_42) == 0) { -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_39, 1); -lean_inc(x_42); -lean_dec(x_39); -x_43 = lean_box(0); -x_44 = l_Lean_Elab_Tactic_Omega_lookup___lambda__2(x_1, x_38, x_38, x_43, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_42); -return x_44; +lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; +x_43 = l_Lean_Elab_Tactic_Omega_lookup___closed__2; +x_44 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_43, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_41); +x_45 = lean_ctor_get(x_44, 0); +lean_inc(x_45); +x_46 = lean_unbox(x_45); +lean_dec(x_45); +if (x_46 == 0) +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_44, 1); +lean_inc(x_47); +lean_dec(x_44); +x_48 = lean_box(0); +x_49 = l_Lean_Elab_Tactic_Omega_lookup___lambda__2(x_40, x_43, x_43, x_48, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_47); +return x_49; } 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; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_45 = lean_ctor_get(x_39, 1); -lean_inc(x_45); -lean_dec(x_39); -lean_inc(x_1); -x_46 = l_Lean_MessageData_ofExpr(x_1); -x_47 = l_Lean_Elab_Tactic_Omega_lookup___closed__4; -x_48 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_48, 0, x_47); -lean_ctor_set(x_48, 1, x_46); -x_49 = l_Lean_Elab_Tactic_Omega_lookup___lambda__2___closed__4; -x_50 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_50, 0, x_48); -lean_ctor_set(x_50, 1, x_49); -x_51 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_38, x_50, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_45); -x_52 = lean_ctor_get(x_51, 0); -lean_inc(x_52); -x_53 = lean_ctor_get(x_51, 1); -lean_inc(x_53); -lean_dec(x_51); -x_54 = l_Lean_Elab_Tactic_Omega_lookup___lambda__2(x_1, x_38, x_38, x_52, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_53); -return x_54; +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_50 = lean_ctor_get(x_44, 1); +lean_inc(x_50); +lean_dec(x_44); +lean_inc(x_40); +x_51 = l_Lean_MessageData_ofExpr(x_40); +x_52 = l_Lean_Elab_Tactic_Omega_lookup___closed__4; +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_Elab_Tactic_Omega_lookup___lambda__2___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 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_43, x_55, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_50); +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_56, 1); +lean_inc(x_58); +lean_dec(x_56); +x_59 = l_Lean_Elab_Tactic_Omega_lookup___lambda__2(x_40, x_43, x_43, x_57, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_58); +return x_59; +} +} +else +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; +lean_dec(x_40); +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); +x_60 = lean_ctor_get(x_42, 0); +lean_inc(x_60); +lean_dec(x_42); +x_61 = lean_box(0); +x_62 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_62, 0, x_60); +lean_ctor_set(x_62, 1, x_61); +x_63 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_63, 0, x_62); +lean_ctor_set(x_63, 1, x_41); +return x_63; +} } } else { -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +uint8_t x_64; +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); -lean_dec(x_1); -x_55 = lean_ctor_get(x_37, 0); -lean_inc(x_55); -lean_dec(x_37); -x_56 = lean_box(0); -x_57 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_57, 0, x_55); -lean_ctor_set(x_57, 1, x_56); -x_58 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_58, 0, x_57); -lean_ctor_set(x_58, 1, x_36); -return x_58; +x_64 = !lean_is_exclusive(x_15); +if (x_64 == 0) +{ +return x_15; +} +else +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_65 = lean_ctor_get(x_15, 0); +x_66 = lean_ctor_get(x_15, 1); +lean_inc(x_66); +lean_inc(x_65); +lean_dec(x_15); +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; } } } @@ -7197,19 +7505,22 @@ lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___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_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___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_10; -x_10 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +uint8_t x_12; lean_object* x_13; +x_12 = lean_unbox(x_5); +lean_dec(x_5); +x_13 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_Omega_lookup___spec__3(x_1, x_2, x_3, x_4, x_12, 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_10; +return x_13; } } LEAN_EXPORT lean_object* l_Lean_AssocList_contains___at_Lean_Elab_Tactic_Omega_lookup___spec__5___boxed(lean_object* x_1, lean_object* x_2) { @@ -7236,46 +7547,75 @@ lean_dec(x_1); return x_7; } } -LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Elab_Tactic_Omega_lookup___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_EXPORT lean_object* l_List_mapM_loop___at_Lean_Elab_Tactic_Omega_lookup___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: { -lean_object* x_11; -x_11 = l_List_mapM_loop___at_Lean_Elab_Tactic_Omega_lookup___spec__13(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +uint8_t x_13; lean_object* x_14; +x_13 = lean_unbox(x_6); +lean_dec(x_6); +x_14 = l_List_mapM_loop___at_Lean_Elab_Tactic_Omega_lookup___spec__13(x_1, x_2, x_3, x_4, x_5, x_13, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_7); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -return x_11; +return x_14; } } -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___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_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___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: { -lean_object* x_11; -x_11 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +uint8_t x_13; lean_object* x_14; +x_13 = lean_unbox(x_6); +lean_dec(x_6); +x_14 = l_Lean_addTrace___at_Lean_Elab_Tactic_Omega_lookup___spec__14(x_1, x_2, x_3, x_4, x_5, x_13, 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); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -return x_11; +return x_14; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_lookup___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_lookup___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_object* x_12; -x_12 = l_Lean_Elab_Tactic_Omega_lookup___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +uint8_t x_14; lean_object* x_15; +x_14 = lean_unbox(x_7); +lean_dec(x_7); +x_15 = l_Lean_Elab_Tactic_Omega_lookup___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_14, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -return x_12; +return x_15; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_lookup___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_8); +lean_dec(x_8); +x_16 = l_Lean_Elab_Tactic_Omega_lookup___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_15, x_9, x_10, x_11, x_12, x_13, x_14); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_lookup___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_5); +lean_dec(x_5); +x_13 = l_Lean_Elab_Tactic_Omega_lookup(x_1, x_2, x_3, x_4, x_12, x_6, x_7, x_8, x_9, x_10, x_11); +return x_13; } } lean_object* initialize_Init_Omega_LinearCombo(uint8_t builtin, lean_object*); @@ -7283,6 +7623,7 @@ lean_object* initialize_Init_Omega_Int(uint8_t builtin, lean_object*); lean_object* initialize_Init_Omega_Logic(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_BitVec_Basic(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_AppBuilder(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Canonicalizer(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Elab_Tactic_Omega_OmegaM(uint8_t builtin, lean_object* w) { lean_object * res; @@ -7303,10 +7644,15 @@ lean_dec_ref(res); res = initialize_Lean_Meta_AppBuilder(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_Meta_Canonicalizer(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); l_Lean_Elab_Tactic_Omega_State_atoms___default = _init_l_Lean_Elab_Tactic_Omega_State_atoms___default(); lean_mark_persistent(l_Lean_Elab_Tactic_Omega_State_atoms___default); l_Lean_Elab_Tactic_Omega_OmegaM_run___rarg___closed__1 = _init_l_Lean_Elab_Tactic_Omega_OmegaM_run___rarg___closed__1(); lean_mark_persistent(l_Lean_Elab_Tactic_Omega_OmegaM_run___rarg___closed__1); +l_Lean_Elab_Tactic_Omega_OmegaM_run___rarg___closed__2 = _init_l_Lean_Elab_Tactic_Omega_OmegaM_run___rarg___closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic_Omega_OmegaM_run___rarg___closed__2); l_Lean_HashMap_toArray___at_Lean_Elab_Tactic_Omega_atoms___spec__1___closed__1 = _init_l_Lean_HashMap_toArray___at_Lean_Elab_Tactic_Omega_atoms___spec__1___closed__1(); lean_mark_persistent(l_Lean_HashMap_toArray___at_Lean_Elab_Tactic_Omega_atoms___spec__1___closed__1); l_Array_qsort_sort___at_Lean_Elab_Tactic_Omega_atoms___spec__4___closed__1 = _init_l_Array_qsort_sort___at_Lean_Elab_Tactic_Omega_atoms___spec__4___closed__1(); diff --git a/stage0/stdlib/Lean/Environment.c b/stage0/stdlib/Lean/Environment.c index 4694d3819ce9..b9a072088401 100644 --- a/stage0/stdlib/Lean/Environment.c +++ b/stage0/stdlib/Lean/Environment.c @@ -20,6 +20,7 @@ LEAN_EXPORT lean_object* l_Lean_ConstantInfo_instantiateTypeLevelParams(lean_obj static lean_object* l___auto____x40_Lean_Environment___hyg_2443____closed__3; LEAN_EXPORT lean_object* l_Lean_throwAlreadyImported___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_displayStats___spec__8___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6351____closed__6; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Environment_find_x3f___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Init_Util_0__outOfBounds___rarg(lean_object*); lean_object* l_Lean_Name_reprPrec(lean_object*, lean_object*); @@ -41,6 +42,7 @@ lean_object* l_Lean_Name_quickLt___boxed(lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_evalConst___boxed(lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_shift_right(size_t, size_t); +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6351____closed__3; LEAN_EXPORT lean_object* l_Lean_ImportStateM_run___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_75____closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_modifyState___rarg___lambda__1(lean_object*, lean_object*); @@ -49,8 +51,10 @@ LEAN_EXPORT uint8_t l_Lean_AssocList_contains___at___private_Lean_Environment_0_ static lean_object* l_Lean_Environment_displayStats___closed__10; LEAN_EXPORT lean_object* l_Lean_registerPersistentEnvExtensionUnsafe(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6351____spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_uint32_to_nat(uint32_t); static lean_object* l_Lean_mkEmptyEnvironment___lambda__1___closed__3; +LEAN_EXPORT uint8_t l___private_Lean_Environment_0__Lean_equivInfo(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_importModules___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4___closed__1; static lean_object* l_Lean_registerPersistentEnvExtensionUnsafe___rarg___closed__2; @@ -74,7 +78,6 @@ static lean_object* l_List_toString___at_Lean_Environment_displayStats___spec__1 LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_mkTagDeclarationExtension___spec__1___boxed(lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_le(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_EnvironmentHeader_regions___default; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MapDeclarationExtension_insert(lean_object*); lean_object* l_Lean_ConstantInfo_levelParams(lean_object*); static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__8; @@ -82,6 +85,7 @@ static lean_object* l_Lean_instInhabitedEnvExtensionInterface___closed__3; static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_75____closed__12; LEAN_EXPORT lean_object* l_Lean_Environment_allImportedModuleNames___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMap_numBuckets___at_Lean_Environment_displayStats___spec__6___boxed(lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6351____closed__1; LEAN_EXPORT lean_object* l_Lean_mkMapDeclarationExtension___rarg___lambda__2___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_setImportedEntries(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -147,6 +151,7 @@ LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_getState(lean_object*, le lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_importModules___lambda__2___closed__1; static lean_object* l_Array_forInUnsafe_loop___at_Lean_importModules___spec__1___closed__2; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6351____spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ConstantInfo_name(lean_object*); LEAN_EXPORT lean_object* l_Lean_TagDeclarationExtension_instInhabitedTagDeclarationExtension; static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_75____closed__16; @@ -164,7 +169,6 @@ static lean_object* l___auto____x40_Lean_Environment___hyg_2443____closed__22; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3(lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMapImp_expand___at_Lean_mkExtNameMap___spec__4(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__4___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_Environment___hyg_6214____lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Environment_imports___boxed(lean_object*); static lean_object* l_List_toString___at_Lean_Environment_displayStats___spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtensionDescr_toArrayFn___default___rarg(lean_object*); @@ -178,9 +182,8 @@ LEAN_EXPORT lean_object* l_Lean_Environment_addDecl___boxed(lean_object*, lean_o static lean_object* l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_ImportStateM_run(lean_object*); LEAN_EXPORT lean_object* l_Lean_importModules___lambda__1(lean_object*, lean_object*, uint32_t, uint8_t, lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__6; uint8_t l_Lean_Name_quickLt(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_HashMap_insertIfNew___at_Lean_finalizeImport___spec__7(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_ConstantInfo_isUnsafe(lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMap_insert___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__6(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__2___spec__1(lean_object*, lean_object*); @@ -200,7 +203,6 @@ LEAN_EXPORT lean_object* l_Lean_updateEnvAttributes___boxed(lean_object*, lean_o LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_importModules___closed__1; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_mkModuleData___spec__4___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_System_FilePath_pathExists(lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Environment___hyg_2443____closed__15; LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___lambda__1(lean_object*); @@ -236,6 +238,7 @@ LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtensionDescr_toArrayFn___de LEAN_EXPORT lean_object* l_Lean_registerEnvExtension(lean_object*); static lean_object* l_Array_qsort_sort___at_Lean_mkMapDeclarationExtension___spec__1___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_mkEmptyEnvironment___spec__1___boxed(lean_object*); +uint8_t l_List_beq___at___private_Lean_Declaration_0__Lean_beqConstantVal____x40_Lean_Declaration___hyg_236____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_mkMapDeclarationExtension___spec__1___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_displayStats___spec__8___lambda__1___closed__3; LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_modifyState___rarg(lean_object*, lean_object*, lean_object*); @@ -268,6 +271,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_mkModuleData___spec__1 uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_getState___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instToStringImport(lean_object*); +uint8_t lean_expr_eqv(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_qsort_sort___at_Lean_mkMapDeclarationExtension___spec__1___rarg___lambda__1(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_registerPersistentEnvExtensionUnsafe___spec__1___rarg(lean_object*, lean_object*, size_t, size_t); static lean_object* l___auto____x40_Lean_Environment___hyg_2443____closed__2; @@ -326,7 +330,6 @@ static size_t l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment static lean_object* l_Lean_throwAlreadyImported___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_MapDeclarationExtension_find_x3f(lean_object*); LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_instInhabitedSimplePersistentEnvExtension___rarg(lean_object*); -LEAN_EXPORT lean_object* l_Lean_HashMap_insert_x27___at_Lean_finalizeImport___spec__7(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_binSearchAux___at_Lean_MapDeclarationExtension_contains___spec__1___rarg(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instMonadEnv___rarg(lean_object*, lean_object*); @@ -349,7 +352,6 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_mkExtNameMap___spe lean_object* lean_list_to_array(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_binSearchAux___at_Lean_MapDeclarationExtension_find_x3f___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withImportModules___rarg(lean_object*, lean_object*, uint32_t, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__4(lean_object*); lean_object* l_EStateM_pure___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtensionState(lean_object*, lean_object*); static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_75____closed__10; @@ -395,6 +397,7 @@ LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_MapDeclarationExtension_fi LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_registerPersistentEnvExtensionUnsafe___boxed(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_initFn____x40_Lean_Environment___hyg_6351____closed__4; LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_mkExtNameMap___spec__6(lean_object*, lean_object*); lean_object* lean_update_env_attributes(lean_object*, lean_object*); @@ -406,6 +409,7 @@ static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_ LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_Environment_isNamespaceName___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_ConstantInfo_instantiateValueLevelParams_x21___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_75____closed__14; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6351____spec__2(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_MapDeclarationExtension_insert___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_instToStringImport___closed__2; static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_75____closed__5; @@ -439,13 +443,13 @@ static lean_object* l_Lean_Environment_displayStats___closed__6; LEAN_EXPORT lean_object* l_Lean_registerEnvExtension___rarg(lean_object*, lean_object*); static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_75____closed__19; LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_Environment_getModuleIdxFor_x3f___spec__2(lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__1; lean_object* lean_usize_to_nat(size_t); size_t lean_hashmap_mk_idx(lean_object*, uint64_t); static lean_object* l_Lean_Environment_evalConstCheck___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__2___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Kernel_isDefEqGuarded___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6351____closed__5; LEAN_EXPORT uint8_t l_Lean_Environment_isNamespace(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvExtension_instInhabitedEnvExtension___boxed(lean_object*, lean_object*); @@ -456,7 +460,6 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_importModules___spec LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Environment_find_x3f___spec__3(lean_object*, size_t, lean_object*); static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_75____closed__2; LEAN_EXPORT lean_object* l_Lean_throwAlreadyImported___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__2(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_mkEmptyEnvironment___spec__2(lean_object*); static lean_object* l_Lean_instInhabitedEnvExtensionInterface___closed__1; LEAN_EXPORT lean_object* l_Lean_HashMapImp_moveEntries___at_Lean_finalizeImport___spec__4(lean_object*, lean_object*, lean_object*); @@ -472,7 +475,6 @@ LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg__ static lean_object* l_Lean_mkTagDeclarationExtension___closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_mkModuleData___spec__5___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_EnvExtensionStateSpec___closed__1; -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6214____lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* lean_environment_add(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Environment_getModuleIdx_x3f___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkModuleData___lambda__2(lean_object*, lean_object*, lean_object*); @@ -494,7 +496,7 @@ LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_setState___rarg___lambda_ static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_75____closed__18; LEAN_EXPORT lean_object* l_Lean_Environment_getModuleIdx_x3f___boxed(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_importModules___spec__1___closed__1; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__3(lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6351____lambda__2(lean_object*); extern lean_object* l_Lean_NameSet_empty; static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_modifyState___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_mkMapDeclarationExtension___rarg___lambda__2(lean_object*); @@ -503,7 +505,6 @@ LEAN_EXPORT lean_object* l_Lean_SMap_stageSizes___at_Lean_Environment_displaySta static lean_object* l_Lean_MapDeclarationExtension_instInhabitedMapDeclarationExtension___closed__1; lean_object* lean_string_length(lean_object*); static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__6; -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__3; static lean_object* l_Lean_instReprImport___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___rarg___lambda__1(lean_object*, lean_object*, lean_object*); @@ -515,6 +516,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_finalizePersistent LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_displayStats___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_finalizeImport___spec__9(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Environment___hyg_2443____closed__19; +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6351____lambda__2___closed__1; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_getState(lean_object*, lean_object*); LEAN_EXPORT lean_object* lean_environment_main_module(lean_object*); @@ -523,10 +525,10 @@ LEAN_EXPORT lean_object* l_Lean_EnvExtension_setState___rarg(lean_object*, lean_ LEAN_EXPORT lean_object* l_Lean_AssocList_replace___at_Lean_finalizeImport___spec__6(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Environment_displayStats___closed__1; static lean_object* l_Lean_instInhabitedPersistentEnvExtension___closed__3; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6351____lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvExtension_getState(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__4(size_t, 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_initFn____x40_Lean_Environment___hyg_6214____closed__4; LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_ensureExtensionsArraySize_loop___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtensionDescr_statsFn___default___boxed(lean_object*, lean_object*); @@ -549,7 +551,6 @@ LEAN_EXPORT lean_object* l_Lean_registerPersistentEnvExtensionUnsafe___rarg___la LEAN_EXPORT uint8_t l_Lean_SMap_contains___at_Lean_Environment_addExtraName___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ConstantInfo_instantiateTypeLevelParams___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_TagDeclarationExtension_tag(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6214_(lean_object*); static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__5; uint8_t l_Lean_NameSet_contains(lean_object*, lean_object*); static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_75____closed__8; @@ -579,7 +580,6 @@ static lean_object* l_Lean_instInhabitedPersistentEnvExtension___closed__5; lean_object* lean_nat_mul(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Kernel_isDefEqGuarded(lean_object*, lean_object*, lean_object*, lean_object*); static uint32_t l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___lambda__1___closed__1; -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__2; LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__1___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkMapDeclarationExtension___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_getState___rarg(lean_object*, lean_object*, lean_object*); @@ -599,6 +599,7 @@ LEAN_EXPORT lean_object* l_panic___at_Lean_EnvExtensionInterfaceUnsafe_modifySta lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_CompactedRegion_free___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_Environment_getModuleIdxFor_x3f___spec__2___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6351_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Kernel_whnf___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_getState___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_getModuleIdxFor_x3f(lean_object*, lean_object*); @@ -660,7 +661,6 @@ 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_object* l_List_redLength___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_Environment_find_x3f___spec__6(lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__5; LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__3___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Environment___hyg_2443____closed__23; static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_75____closed__22; @@ -673,11 +673,14 @@ LEAN_EXPORT lean_object* l_Lean_SMap_stageSizes___at_Lean_Environment_displaySta static lean_object* l_Lean_mkModuleData___closed__2; LEAN_EXPORT lean_object* l_Lean_Environment_getNamespaceSet___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMapImp_find_x3f___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6351____spec__5(lean_object*, size_t, size_t, lean_object*); static lean_object* l___auto____x40_Lean_Environment___hyg_2443____closed__24; static lean_object* l_Lean_mkEmptyEnvironment___lambda__1___closed__1; lean_object* lean_string_append(lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Environment___hyg_2443____closed__1; LEAN_EXPORT lean_object* l_Lean_instMonadEnv(lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6351____closed__2; +LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6351____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_evalConstCheck___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_binSearchAux___at_Lean_MapDeclarationExtension_find_x3f___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -697,6 +700,7 @@ lean_object* l_Lean_SMap_insert___at_Lean_NameSSet_insert___spec__1(lean_object* uint8_t lean_usize_dec_lt(size_t, size_t); LEAN_EXPORT lean_object* lean_mk_empty_environment(uint32_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkTagDeclarationExtension___lambda__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6351____spec__4(lean_object*); static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___closed__2; LEAN_EXPORT lean_object* l_Lean_mkMapDeclarationExtension___boxed(lean_object*, lean_object*); static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_getState___rarg___closed__2; @@ -711,11 +715,12 @@ LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_setState(lean_objec lean_object* l_EStateM_bind___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkTagDeclarationExtension(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvironmentHeader_moduleData___default; +LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_equivInfo___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6351____spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* lean_import_modules(lean_object*, lean_object*, uint32_t, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMapImp_moveEntries___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__9(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_registerExt(lean_object*); LEAN_EXPORT lean_object* l_Lean_importModules___boxed__const__1; -LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__1(lean_object*, lean_object*); lean_object* lean_kernel_whnf(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_mkEmptyEnvironment___lambda__1___closed__2; @@ -738,8 +743,8 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at_Lean_Environment static lean_object* l_List_foldl___at_Lean_Environment_displayStats___spec__2___closed__1; static lean_object* l_Lean_instInhabitedPersistentEnvExtension___closed__4; LEAN_EXPORT lean_object* l_Lean_SMap_numBuckets___at_Lean_Environment_displayStats___spec__5___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6351____spec__3(lean_object*, size_t, size_t, lean_object*); lean_object* l_Nat_repr(lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__5(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Environment_addExtraName___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvExtension_setState(lean_object*); LEAN_EXPORT lean_object* l_Lean_withImportModules___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -759,7 +764,6 @@ static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__3; LEAN_EXPORT lean_object* l_Lean_ModuleIdx_toNat___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_finalizeImport___spec__8___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedModuleData___closed__2; -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6214____lambda__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwAlreadyImported(lean_object*); uint8_t l_Array_isEmpty___rarg(lean_object*); static lean_object* _init_l_Lean_EnvExtensionStateSpec___closed__1() { @@ -11463,6 +11467,88 @@ lean_dec(x_1); return x_4; } } +LEAN_EXPORT uint8_t l___private_Lean_Environment_0__Lean_equivInfo(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 2) +{ +if (lean_obj_tag(x_2) == 2) +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_3 = lean_ctor_get(x_1, 0); +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_3, 0); +x_6 = lean_ctor_get(x_5, 0); +x_7 = lean_ctor_get(x_4, 0); +x_8 = lean_ctor_get(x_7, 0); +x_9 = lean_name_eq(x_6, x_8); +if (x_9 == 0) +{ +uint8_t x_10; +x_10 = 0; +return x_10; +} +else +{ +lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_11 = lean_ctor_get(x_5, 2); +x_12 = lean_ctor_get(x_7, 2); +x_13 = lean_expr_eqv(x_11, x_12); +if (x_13 == 0) +{ +uint8_t x_14; +x_14 = 0; +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_15 = lean_ctor_get(x_5, 1); +x_16 = lean_ctor_get(x_7, 1); +x_17 = l_List_beq___at___private_Lean_Declaration_0__Lean_beqConstantVal____x40_Lean_Declaration___hyg_236____spec__1(x_15, x_16); +if (x_17 == 0) +{ +uint8_t x_18; +x_18 = 0; +return x_18; +} +else +{ +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = lean_ctor_get(x_3, 2); +x_20 = lean_ctor_get(x_4, 2); +x_21 = l_List_beq___at___private_Lean_Declaration_0__Lean_beqConstantVal____x40_Lean_Declaration___hyg_236____spec__1(x_19, x_20); +return x_21; +} +} +} +} +else +{ +uint8_t x_22; +x_22 = 0; +return x_22; +} +} +else +{ +uint8_t x_23; +x_23 = 0; +return x_23; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_equivInfo___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Lean_Environment_0__Lean_equivInfo(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} LEAN_EXPORT uint8_t l_Lean_AssocList_contains___at_Lean_finalizeImport___spec__2(lean_object* x_1, lean_object* x_2) { _start: { @@ -11779,56 +11865,60 @@ return x_38; } } } -LEAN_EXPORT lean_object* l_Lean_HashMap_insert_x27___at_Lean_finalizeImport___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_HashMap_insertIfNew___at_Lean_finalizeImport___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_4; lean_object* x_5; lean_object* x_6; uint64_t x_7; size_t x_8; lean_object* x_9; lean_object* x_10; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +x_5 = lean_ctor_get(x_1, 1); +lean_inc(x_5); +x_6 = lean_array_get_size(x_5); +x_7 = l_Lean_Name_hash___override(x_2); +lean_inc(x_6); +x_8 = lean_hashmap_mk_idx(x_6, x_7); +x_9 = lean_array_uget(x_5, x_8); +x_10 = l_Lean_AssocList_find_x3f___at_Lean_Environment_find_x3f___spec__6(x_2, x_9); +if (lean_obj_tag(x_10) == 0) { -lean_object* x_5; lean_object* x_6; lean_object* x_7; uint64_t x_8; size_t x_9; lean_object* x_10; uint8_t x_11; -x_5 = lean_ctor_get(x_1, 0); -x_6 = lean_ctor_get(x_1, 1); -x_7 = lean_array_get_size(x_6); -x_8 = l_Lean_Name_hash___override(x_2); -lean_inc(x_7); -x_9 = lean_hashmap_mk_idx(x_7, x_8); -x_10 = lean_array_uget(x_6, x_9); -x_11 = l_Lean_AssocList_contains___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__7(x_2, x_10); +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_unsigned_to_nat(1u); -x_13 = lean_nat_add(x_5, x_12); -lean_dec(x_5); -x_14 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_14, 0, x_2); -lean_ctor_set(x_14, 1, x_3); -lean_ctor_set(x_14, 2, x_10); -x_15 = lean_array_uset(x_6, x_9, x_14); -x_16 = l___private_Lean_Data_HashMap_0__Lean_numBucketsForCapacity(x_13); -x_17 = lean_nat_dec_le(x_16, x_7); -lean_dec(x_7); -lean_dec(x_16); -if (x_17 == 0) +lean_object* x_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_12 = lean_ctor_get(x_1, 1); +lean_dec(x_12); +x_13 = lean_ctor_get(x_1, 0); +lean_dec(x_13); +x_14 = lean_unsigned_to_nat(1u); +x_15 = lean_nat_add(x_4, x_14); +lean_dec(x_4); +x_16 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_16, 0, x_2); +lean_ctor_set(x_16, 1, x_3); +lean_ctor_set(x_16, 2, x_9); +x_17 = lean_array_uset(x_5, x_8, x_16); +x_18 = l___private_Lean_Data_HashMap_0__Lean_numBucketsForCapacity(x_15); +x_19 = lean_nat_dec_le(x_18, x_6); +lean_dec(x_6); +lean_dec(x_18); +if (x_19 == 0) { -lean_object* x_18; uint8_t x_19; lean_object* x_20; lean_object* x_21; +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_free_object(x_1); -x_18 = l_Lean_HashMapImp_expand___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__8(x_13, x_15); -x_19 = 0; -x_20 = lean_box(x_19); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_18); -lean_ctor_set(x_21, 1, x_20); -return x_21; +x_20 = l_Lean_HashMapImp_expand___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__8(x_15, x_17); +x_21 = lean_box(0); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +return x_22; } else { -uint8_t x_22; lean_object* x_23; lean_object* x_24; -lean_ctor_set(x_1, 1, x_15); -lean_ctor_set(x_1, 0, x_13); -x_22 = 0; -x_23 = lean_box(x_22); +lean_object* x_23; lean_object* x_24; +lean_ctor_set(x_1, 1, x_17); +lean_ctor_set(x_1, 0, x_15); +x_23 = lean_box(0); x_24 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_24, 0, x_1); lean_ctor_set(x_24, 1, x_23); @@ -11837,88 +11927,74 @@ return x_24; } else { -lean_object* x_25; lean_object* x_26; uint8_t x_27; lean_object* x_28; lean_object* x_29; -lean_dec(x_7); -x_25 = l_Lean_AssocList_replace___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__11(x_2, x_3, x_10); -x_26 = lean_array_uset(x_6, x_9, x_25); -lean_ctor_set(x_1, 1, x_26); -x_27 = 1; -x_28 = lean_box(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_28); -return x_29; -} -} -else -{ -lean_object* x_30; lean_object* x_31; lean_object* x_32; uint64_t x_33; size_t x_34; lean_object* x_35; uint8_t x_36; -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_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; lean_dec(x_1); -x_32 = lean_array_get_size(x_31); -x_33 = l_Lean_Name_hash___override(x_2); -lean_inc(x_32); -x_34 = lean_hashmap_mk_idx(x_32, x_33); -x_35 = lean_array_uget(x_31, x_34); -x_36 = l_Lean_AssocList_contains___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__7(x_2, x_35); -if (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; -x_37 = lean_unsigned_to_nat(1u); -x_38 = lean_nat_add(x_30, x_37); -lean_dec(x_30); -x_39 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_39, 0, x_2); -lean_ctor_set(x_39, 1, x_3); -lean_ctor_set(x_39, 2, x_35); -x_40 = lean_array_uset(x_31, x_34, x_39); -x_41 = l___private_Lean_Data_HashMap_0__Lean_numBucketsForCapacity(x_38); -x_42 = lean_nat_dec_le(x_41, x_32); -lean_dec(x_32); -lean_dec(x_41); -if (x_42 == 0) +x_25 = lean_unsigned_to_nat(1u); +x_26 = lean_nat_add(x_4, x_25); +lean_dec(x_4); +x_27 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_27, 0, x_2); +lean_ctor_set(x_27, 1, x_3); +lean_ctor_set(x_27, 2, x_9); +x_28 = lean_array_uset(x_5, x_8, x_27); +x_29 = l___private_Lean_Data_HashMap_0__Lean_numBucketsForCapacity(x_26); +x_30 = lean_nat_dec_le(x_29, x_6); +lean_dec(x_6); +lean_dec(x_29); +if (x_30 == 0) { -lean_object* x_43; uint8_t x_44; lean_object* x_45; lean_object* x_46; -x_43 = l_Lean_HashMapImp_expand___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__8(x_38, x_40); -x_44 = 0; -x_45 = lean_box(x_44); -x_46 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_46, 0, x_43); -lean_ctor_set(x_46, 1, x_45); -return x_46; +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = l_Lean_HashMapImp_expand___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__8(x_26, x_28); +x_32 = lean_box(0); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; } else { -lean_object* x_47; uint8_t x_48; lean_object* x_49; lean_object* x_50; -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_38); -lean_ctor_set(x_47, 1, x_40); -x_48 = 0; -x_49 = lean_box(x_48); -x_50 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_50, 0, x_47); -lean_ctor_set(x_50, 1, x_49); -return x_50; +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_26); +lean_ctor_set(x_34, 1, x_28); +x_35 = lean_box(0); +x_36 = lean_alloc_ctor(0, 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_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; lean_object* x_55; lean_object* x_56; -lean_dec(x_32); -x_51 = l_Lean_AssocList_replace___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__11(x_2, x_3, x_35); -x_52 = lean_array_uset(x_31, x_34, x_51); -x_53 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_53, 0, x_30); -lean_ctor_set(x_53, 1, x_52); -x_54 = 1; -x_55 = lean_box(x_54); -x_56 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_56, 0, x_53); -lean_ctor_set(x_56, 1, x_55); -return x_56; +uint8_t x_37; +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_37 = !lean_is_exclusive(x_10); +if (x_37 == 0) +{ +lean_object* x_38; +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_1); +lean_ctor_set(x_38, 1, x_10); +return x_38; +} +else +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_10, 0); +lean_inc(x_39); +lean_dec(x_10); +x_40 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_40, 0, x_39); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_1); +lean_ctor_set(x_41, 1, x_40); +return x_41; } } } @@ -12000,7 +12076,7 @@ lean_free_object(x_6); x_22 = !lean_is_exclusive(x_14); 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; uint8_t x_31; +lean_object* x_23; lean_object* x_24; 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_23 = lean_ctor_get(x_14, 2); lean_dec(x_23); x_24 = lean_ctor_get(x_14, 1); @@ -12012,387 +12088,523 @@ x_27 = lean_unsigned_to_nat(1u); x_28 = lean_nat_add(x_18, x_27); lean_dec(x_18); lean_ctor_set(x_14, 1, x_28); +lean_inc(x_26); lean_inc(x_10); -x_29 = l_Lean_HashMap_insert_x27___at_Lean_finalizeImport___spec__7(x_16, x_10, x_26); +x_29 = l_Lean_HashMap_insertIfNew___at_Lean_finalizeImport___spec__7(x_16, x_10, x_26); x_30 = lean_ctor_get(x_29, 1); lean_inc(x_30); -x_31 = lean_unbox(x_30); -lean_dec(x_30); -if (x_31 == 0) +if (lean_obj_tag(x_30) == 0) { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; size_t x_38; size_t x_39; -x_32 = lean_ctor_get(x_29, 0); -lean_inc(x_32); +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; size_t x_37; size_t x_38; +lean_dec(x_26); +x_31 = lean_ctor_get(x_29, 0); +lean_inc(x_31); lean_dec(x_29); -x_33 = lean_box(0); +x_32 = lean_box(0); lean_inc(x_2); -x_34 = l_Array_forInUnsafe_loop___at_Lean_finalizeImport___spec__8___lambda__1(x_10, x_2, x_14, x_15, x_32, x_33, x_7); -x_35 = lean_ctor_get(x_34, 0); +x_33 = l_Array_forInUnsafe_loop___at_Lean_finalizeImport___spec__8___lambda__1(x_10, x_2, x_14, x_15, x_31, x_32, x_7); +x_34 = lean_ctor_get(x_33, 0); +lean_inc(x_34); +x_35 = lean_ctor_get(x_33, 1); lean_inc(x_35); -x_36 = lean_ctor_get(x_34, 1); +lean_dec(x_33); +x_36 = lean_ctor_get(x_34, 0); lean_inc(x_36); lean_dec(x_34); -x_37 = lean_ctor_get(x_35, 0); -lean_inc(x_37); -lean_dec(x_35); -x_38 = 1; -x_39 = lean_usize_add(x_5, x_38); -x_5 = x_39; -x_6 = x_37; -x_7 = x_36; +x_37 = 1; +x_38 = lean_usize_add(x_5, x_37); +x_5 = x_38; +x_6 = x_36; +x_7 = x_35; goto _start; } else { -lean_object* x_41; uint8_t x_42; +lean_object* x_40; lean_object* x_41; uint8_t x_42; +x_40 = lean_ctor_get(x_29, 0); +lean_inc(x_40); lean_dec(x_29); +x_41 = lean_ctor_get(x_30, 0); +lean_inc(x_41); +lean_dec(x_30); +x_42 = l___private_Lean_Environment_0__Lean_equivInfo(x_41, x_26); +lean_dec(x_26); +lean_dec(x_41); +if (x_42 == 0) +{ +lean_object* x_43; uint8_t x_44; +lean_dec(x_40); lean_dec(x_14); -x_41 = l_Lean_throwAlreadyImported___rarg(x_1, x_15, x_2, x_10, x_7); +x_43 = l_Lean_throwAlreadyImported___rarg(x_1, x_15, x_2, x_10, x_7); lean_dec(x_2); -x_42 = !lean_is_exclusive(x_41); -if (x_42 == 0) +x_44 = !lean_is_exclusive(x_43); +if (x_44 == 0) { -return x_41; +return x_43; } else { -lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_43 = lean_ctor_get(x_41, 0); -x_44 = lean_ctor_get(x_41, 1); -lean_inc(x_44); -lean_inc(x_43); -lean_dec(x_41); -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_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_43, 0); +x_46 = lean_ctor_get(x_43, 1); +lean_inc(x_46); +lean_inc(x_45); +lean_dec(x_43); +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_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; size_t x_53; size_t x_54; +x_48 = lean_box(0); +lean_inc(x_2); +x_49 = l_Array_forInUnsafe_loop___at_Lean_finalizeImport___spec__8___lambda__1(x_10, x_2, x_14, x_15, x_40, x_48, x_7); +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_50, 0); +lean_inc(x_52); +lean_dec(x_50); +x_53 = 1; +x_54 = lean_usize_add(x_5, x_53); +x_5 = x_54; +x_6 = x_52; +x_7 = x_51; +goto _start; } } } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; +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_14); -x_46 = lean_array_fget(x_17, x_18); -x_47 = lean_unsigned_to_nat(1u); -x_48 = lean_nat_add(x_18, x_47); +x_56 = lean_array_fget(x_17, x_18); +x_57 = lean_unsigned_to_nat(1u); +x_58 = lean_nat_add(x_18, x_57); lean_dec(x_18); -x_49 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_49, 0, x_17); -lean_ctor_set(x_49, 1, x_48); -lean_ctor_set(x_49, 2, x_19); +x_59 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_59, 0, x_17); +lean_ctor_set(x_59, 1, x_58); +lean_ctor_set(x_59, 2, x_19); +lean_inc(x_56); lean_inc(x_10); -x_50 = l_Lean_HashMap_insert_x27___at_Lean_finalizeImport___spec__7(x_16, x_10, x_46); -x_51 = lean_ctor_get(x_50, 1); -lean_inc(x_51); -x_52 = lean_unbox(x_51); -lean_dec(x_51); -if (x_52 == 0) +x_60 = l_Lean_HashMap_insertIfNew___at_Lean_finalizeImport___spec__7(x_16, x_10, x_56); +x_61 = lean_ctor_get(x_60, 1); +lean_inc(x_61); +if (lean_obj_tag(x_61) == 0) { -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; size_t x_59; size_t x_60; -x_53 = lean_ctor_get(x_50, 0); -lean_inc(x_53); -lean_dec(x_50); -x_54 = lean_box(0); -lean_inc(x_2); -x_55 = l_Array_forInUnsafe_loop___at_Lean_finalizeImport___spec__8___lambda__1(x_10, x_2, x_49, x_15, x_53, x_54, x_7); -x_56 = lean_ctor_get(x_55, 0); -lean_inc(x_56); -x_57 = lean_ctor_get(x_55, 1); -lean_inc(x_57); -lean_dec(x_55); -x_58 = lean_ctor_get(x_56, 0); -lean_inc(x_58); +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; size_t x_68; size_t x_69; lean_dec(x_56); -x_59 = 1; -x_60 = lean_usize_add(x_5, x_59); -x_5 = x_60; -x_6 = x_58; -x_7 = x_57; +x_62 = lean_ctor_get(x_60, 0); +lean_inc(x_62); +lean_dec(x_60); +x_63 = lean_box(0); +lean_inc(x_2); +x_64 = l_Array_forInUnsafe_loop___at_Lean_finalizeImport___spec__8___lambda__1(x_10, x_2, x_59, x_15, x_62, x_63, x_7); +x_65 = lean_ctor_get(x_64, 0); +lean_inc(x_65); +x_66 = lean_ctor_get(x_64, 1); +lean_inc(x_66); +lean_dec(x_64); +x_67 = lean_ctor_get(x_65, 0); +lean_inc(x_67); +lean_dec(x_65); +x_68 = 1; +x_69 = lean_usize_add(x_5, x_68); +x_5 = x_69; +x_6 = x_67; +x_7 = x_66; goto _start; } else { -lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; -lean_dec(x_50); -lean_dec(x_49); -x_62 = l_Lean_throwAlreadyImported___rarg(x_1, x_15, x_2, x_10, x_7); +lean_object* x_71; lean_object* x_72; uint8_t x_73; +x_71 = lean_ctor_get(x_60, 0); +lean_inc(x_71); +lean_dec(x_60); +x_72 = lean_ctor_get(x_61, 0); +lean_inc(x_72); +lean_dec(x_61); +x_73 = l___private_Lean_Environment_0__Lean_equivInfo(x_72, x_56); +lean_dec(x_56); +lean_dec(x_72); +if (x_73 == 0) +{ +lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_dec(x_71); +lean_dec(x_59); +x_74 = l_Lean_throwAlreadyImported___rarg(x_1, x_15, x_2, x_10, x_7); lean_dec(x_2); -x_63 = lean_ctor_get(x_62, 0); -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; +x_75 = lean_ctor_get(x_74, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_74, 1); +lean_inc(x_76); +if (lean_is_exclusive(x_74)) { + lean_ctor_release(x_74, 0); + lean_ctor_release(x_74, 1); + x_77 = x_74; } else { - lean_dec_ref(x_62); - x_65 = lean_box(0); + lean_dec_ref(x_74); + x_77 = lean_box(0); } -if (lean_is_scalar(x_65)) { - x_66 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_77)) { + x_78 = lean_alloc_ctor(1, 2, 0); } else { - x_66 = x_65; + 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; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; size_t x_84; size_t x_85; +x_79 = lean_box(0); +lean_inc(x_2); +x_80 = l_Array_forInUnsafe_loop___at_Lean_finalizeImport___spec__8___lambda__1(x_10, x_2, x_59, x_15, x_71, x_79, x_7); +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 = 1; +x_85 = lean_usize_add(x_5, x_84); +x_5 = x_85; +x_6 = x_83; +x_7 = x_82; +goto _start; } -lean_ctor_set(x_66, 0, x_63); -lean_ctor_set(x_66, 1, x_64); -return x_66; } } } } else { -lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; uint8_t x_73; -x_67 = lean_ctor_get(x_6, 0); -x_68 = lean_ctor_get(x_12, 0); -x_69 = lean_ctor_get(x_12, 1); -lean_inc(x_69); -lean_inc(x_68); +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_87 = lean_ctor_get(x_6, 0); +x_88 = lean_ctor_get(x_12, 0); +x_89 = lean_ctor_get(x_12, 1); +lean_inc(x_89); +lean_inc(x_88); lean_dec(x_12); -x_70 = lean_ctor_get(x_67, 0); -lean_inc(x_70); -x_71 = lean_ctor_get(x_67, 1); -lean_inc(x_71); -x_72 = lean_ctor_get(x_67, 2); -lean_inc(x_72); -x_73 = lean_nat_dec_lt(x_71, x_72); -if (x_73 == 0) +x_90 = lean_ctor_get(x_87, 0); +lean_inc(x_90); +x_91 = lean_ctor_get(x_87, 1); +lean_inc(x_91); +x_92 = lean_ctor_get(x_87, 2); +lean_inc(x_92); +x_93 = lean_nat_dec_lt(x_91, x_92); +if (x_93 == 0) { -lean_object* x_74; lean_object* x_75; -lean_dec(x_72); -lean_dec(x_71); -lean_dec(x_70); +lean_object* x_94; lean_object* x_95; +lean_dec(x_92); +lean_dec(x_91); +lean_dec(x_90); lean_dec(x_10); lean_dec(x_2); -x_74 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_74, 0, x_68); -lean_ctor_set(x_74, 1, x_69); -lean_ctor_set(x_6, 1, x_74); -x_75 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_75, 0, x_6); -lean_ctor_set(x_75, 1, x_7); -return x_75; +x_94 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_94, 0, x_88); +lean_ctor_set(x_94, 1, x_89); +lean_ctor_set(x_6, 1, x_94); +x_95 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_95, 0, x_6); +lean_ctor_set(x_95, 1, x_7); +return x_95; } 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; uint8_t x_83; +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_free_object(x_6); -if (lean_is_exclusive(x_67)) { - lean_ctor_release(x_67, 0); - lean_ctor_release(x_67, 1); - lean_ctor_release(x_67, 2); - x_76 = x_67; +if (lean_is_exclusive(x_87)) { + lean_ctor_release(x_87, 0); + lean_ctor_release(x_87, 1); + lean_ctor_release(x_87, 2); + x_96 = x_87; } else { - lean_dec_ref(x_67); - x_76 = lean_box(0); + lean_dec_ref(x_87); + x_96 = lean_box(0); } -x_77 = lean_array_fget(x_70, x_71); -x_78 = lean_unsigned_to_nat(1u); -x_79 = lean_nat_add(x_71, x_78); -lean_dec(x_71); -if (lean_is_scalar(x_76)) { - x_80 = lean_alloc_ctor(0, 3, 0); +x_97 = lean_array_fget(x_90, x_91); +x_98 = lean_unsigned_to_nat(1u); +x_99 = lean_nat_add(x_91, x_98); +lean_dec(x_91); +if (lean_is_scalar(x_96)) { + x_100 = lean_alloc_ctor(0, 3, 0); } else { - x_80 = x_76; + x_100 = x_96; } -lean_ctor_set(x_80, 0, x_70); -lean_ctor_set(x_80, 1, x_79); -lean_ctor_set(x_80, 2, x_72); +lean_ctor_set(x_100, 0, x_90); +lean_ctor_set(x_100, 1, x_99); +lean_ctor_set(x_100, 2, x_92); +lean_inc(x_97); lean_inc(x_10); -x_81 = l_Lean_HashMap_insert_x27___at_Lean_finalizeImport___spec__7(x_69, x_10, x_77); -x_82 = lean_ctor_get(x_81, 1); -lean_inc(x_82); -x_83 = lean_unbox(x_82); -lean_dec(x_82); -if (x_83 == 0) +x_101 = l_Lean_HashMap_insertIfNew___at_Lean_finalizeImport___spec__7(x_89, x_10, x_97); +x_102 = lean_ctor_get(x_101, 1); +lean_inc(x_102); +if (lean_obj_tag(x_102) == 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; size_t x_90; size_t x_91; -x_84 = lean_ctor_get(x_81, 0); -lean_inc(x_84); -lean_dec(x_81); -x_85 = lean_box(0); +lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; size_t x_109; size_t x_110; +lean_dec(x_97); +x_103 = lean_ctor_get(x_101, 0); +lean_inc(x_103); +lean_dec(x_101); +x_104 = lean_box(0); lean_inc(x_2); -x_86 = l_Array_forInUnsafe_loop___at_Lean_finalizeImport___spec__8___lambda__1(x_10, x_2, x_80, x_68, x_84, x_85, x_7); -x_87 = lean_ctor_get(x_86, 0); -lean_inc(x_87); -x_88 = lean_ctor_get(x_86, 1); -lean_inc(x_88); -lean_dec(x_86); -x_89 = lean_ctor_get(x_87, 0); -lean_inc(x_89); -lean_dec(x_87); -x_90 = 1; -x_91 = lean_usize_add(x_5, x_90); -x_5 = x_91; -x_6 = x_89; -x_7 = x_88; +x_105 = l_Array_forInUnsafe_loop___at_Lean_finalizeImport___spec__8___lambda__1(x_10, x_2, x_100, x_88, x_103, x_104, x_7); +x_106 = lean_ctor_get(x_105, 0); +lean_inc(x_106); +x_107 = lean_ctor_get(x_105, 1); +lean_inc(x_107); +lean_dec(x_105); +x_108 = lean_ctor_get(x_106, 0); +lean_inc(x_108); +lean_dec(x_106); +x_109 = 1; +x_110 = lean_usize_add(x_5, x_109); +x_5 = x_110; +x_6 = x_108; +x_7 = x_107; goto _start; } else { -lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; -lean_dec(x_81); -lean_dec(x_80); -x_93 = l_Lean_throwAlreadyImported___rarg(x_1, x_68, x_2, x_10, x_7); +lean_object* x_112; lean_object* x_113; uint8_t x_114; +x_112 = lean_ctor_get(x_101, 0); +lean_inc(x_112); +lean_dec(x_101); +x_113 = lean_ctor_get(x_102, 0); +lean_inc(x_113); +lean_dec(x_102); +x_114 = l___private_Lean_Environment_0__Lean_equivInfo(x_113, x_97); +lean_dec(x_97); +lean_dec(x_113); +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_dec(x_112); +lean_dec(x_100); +x_115 = l_Lean_throwAlreadyImported___rarg(x_1, x_88, x_2, x_10, x_7); lean_dec(x_2); -x_94 = lean_ctor_get(x_93, 0); -lean_inc(x_94); -x_95 = lean_ctor_get(x_93, 1); -lean_inc(x_95); -if (lean_is_exclusive(x_93)) { - lean_ctor_release(x_93, 0); - lean_ctor_release(x_93, 1); - x_96 = x_93; +x_116 = lean_ctor_get(x_115, 0); +lean_inc(x_116); +x_117 = lean_ctor_get(x_115, 1); +lean_inc(x_117); +if (lean_is_exclusive(x_115)) { + lean_ctor_release(x_115, 0); + lean_ctor_release(x_115, 1); + x_118 = x_115; } else { - lean_dec_ref(x_93); - x_96 = lean_box(0); + lean_dec_ref(x_115); + x_118 = lean_box(0); } -if (lean_is_scalar(x_96)) { - x_97 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_118)) { + x_119 = lean_alloc_ctor(1, 2, 0); } else { - x_97 = x_96; + x_119 = x_118; +} +lean_ctor_set(x_119, 0, x_116); +lean_ctor_set(x_119, 1, x_117); +return x_119; +} +else +{ +lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; size_t x_125; size_t x_126; +x_120 = lean_box(0); +lean_inc(x_2); +x_121 = l_Array_forInUnsafe_loop___at_Lean_finalizeImport___spec__8___lambda__1(x_10, x_2, x_100, x_88, x_112, x_120, x_7); +x_122 = lean_ctor_get(x_121, 0); +lean_inc(x_122); +x_123 = lean_ctor_get(x_121, 1); +lean_inc(x_123); +lean_dec(x_121); +x_124 = lean_ctor_get(x_122, 0); +lean_inc(x_124); +lean_dec(x_122); +x_125 = 1; +x_126 = lean_usize_add(x_5, x_125); +x_5 = x_126; +x_6 = x_124; +x_7 = x_123; +goto _start; } -lean_ctor_set(x_97, 0, x_94); -lean_ctor_set(x_97, 1, x_95); -return x_97; } } } } 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; -x_98 = lean_ctor_get(x_6, 1); -x_99 = lean_ctor_get(x_6, 0); -lean_inc(x_98); -lean_inc(x_99); +lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; uint8_t x_136; +x_128 = lean_ctor_get(x_6, 1); +x_129 = lean_ctor_get(x_6, 0); +lean_inc(x_128); +lean_inc(x_129); lean_dec(x_6); -x_100 = lean_ctor_get(x_98, 0); -lean_inc(x_100); -x_101 = lean_ctor_get(x_98, 1); -lean_inc(x_101); -if (lean_is_exclusive(x_98)) { - lean_ctor_release(x_98, 0); - lean_ctor_release(x_98, 1); - x_102 = x_98; +x_130 = lean_ctor_get(x_128, 0); +lean_inc(x_130); +x_131 = lean_ctor_get(x_128, 1); +lean_inc(x_131); +if (lean_is_exclusive(x_128)) { + lean_ctor_release(x_128, 0); + lean_ctor_release(x_128, 1); + x_132 = x_128; } else { - lean_dec_ref(x_98); - x_102 = lean_box(0); + lean_dec_ref(x_128); + x_132 = lean_box(0); } -x_103 = lean_ctor_get(x_99, 0); -lean_inc(x_103); -x_104 = lean_ctor_get(x_99, 1); -lean_inc(x_104); -x_105 = lean_ctor_get(x_99, 2); -lean_inc(x_105); -x_106 = lean_nat_dec_lt(x_104, x_105); -if (x_106 == 0) +x_133 = lean_ctor_get(x_129, 0); +lean_inc(x_133); +x_134 = lean_ctor_get(x_129, 1); +lean_inc(x_134); +x_135 = lean_ctor_get(x_129, 2); +lean_inc(x_135); +x_136 = lean_nat_dec_lt(x_134, x_135); +if (x_136 == 0) { -lean_object* x_107; lean_object* x_108; lean_object* x_109; -lean_dec(x_105); -lean_dec(x_104); -lean_dec(x_103); +lean_object* x_137; lean_object* x_138; lean_object* x_139; +lean_dec(x_135); +lean_dec(x_134); +lean_dec(x_133); lean_dec(x_10); lean_dec(x_2); -if (lean_is_scalar(x_102)) { - x_107 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_132)) { + x_137 = lean_alloc_ctor(0, 2, 0); } else { - x_107 = x_102; -} -lean_ctor_set(x_107, 0, x_100); -lean_ctor_set(x_107, 1, x_101); -x_108 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_108, 0, x_99); -lean_ctor_set(x_108, 1, x_107); -x_109 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_109, 0, x_108); -lean_ctor_set(x_109, 1, x_7); -return x_109; + x_137 = x_132; +} +lean_ctor_set(x_137, 0, x_130); +lean_ctor_set(x_137, 1, x_131); +x_138 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_138, 0, x_129); +lean_ctor_set(x_138, 1, x_137); +x_139 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_139, 0, x_138); +lean_ctor_set(x_139, 1, x_7); +return x_139; } 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; uint8_t x_117; -lean_dec(x_102); -if (lean_is_exclusive(x_99)) { - lean_ctor_release(x_99, 0); - lean_ctor_release(x_99, 1); - lean_ctor_release(x_99, 2); - x_110 = x_99; +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_132); +if (lean_is_exclusive(x_129)) { + lean_ctor_release(x_129, 0); + lean_ctor_release(x_129, 1); + lean_ctor_release(x_129, 2); + x_140 = x_129; } else { - lean_dec_ref(x_99); - x_110 = lean_box(0); -} -x_111 = lean_array_fget(x_103, x_104); -x_112 = lean_unsigned_to_nat(1u); -x_113 = lean_nat_add(x_104, x_112); -lean_dec(x_104); -if (lean_is_scalar(x_110)) { - x_114 = lean_alloc_ctor(0, 3, 0); + lean_dec_ref(x_129); + x_140 = lean_box(0); +} +x_141 = lean_array_fget(x_133, x_134); +x_142 = lean_unsigned_to_nat(1u); +x_143 = lean_nat_add(x_134, x_142); +lean_dec(x_134); +if (lean_is_scalar(x_140)) { + x_144 = lean_alloc_ctor(0, 3, 0); } else { - x_114 = x_110; + x_144 = x_140; } -lean_ctor_set(x_114, 0, x_103); -lean_ctor_set(x_114, 1, x_113); -lean_ctor_set(x_114, 2, x_105); +lean_ctor_set(x_144, 0, x_133); +lean_ctor_set(x_144, 1, x_143); +lean_ctor_set(x_144, 2, x_135); +lean_inc(x_141); lean_inc(x_10); -x_115 = l_Lean_HashMap_insert_x27___at_Lean_finalizeImport___spec__7(x_101, x_10, x_111); -x_116 = lean_ctor_get(x_115, 1); -lean_inc(x_116); -x_117 = lean_unbox(x_116); -lean_dec(x_116); -if (x_117 == 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; size_t x_124; size_t x_125; -x_118 = lean_ctor_get(x_115, 0); -lean_inc(x_118); -lean_dec(x_115); -x_119 = lean_box(0); +x_145 = l_Lean_HashMap_insertIfNew___at_Lean_finalizeImport___spec__7(x_131, x_10, x_141); +x_146 = lean_ctor_get(x_145, 1); +lean_inc(x_146); +if (lean_obj_tag(x_146) == 0) +{ +lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; size_t x_153; size_t x_154; +lean_dec(x_141); +x_147 = lean_ctor_get(x_145, 0); +lean_inc(x_147); +lean_dec(x_145); +x_148 = lean_box(0); lean_inc(x_2); -x_120 = l_Array_forInUnsafe_loop___at_Lean_finalizeImport___spec__8___lambda__1(x_10, x_2, x_114, x_100, x_118, x_119, x_7); -x_121 = lean_ctor_get(x_120, 0); -lean_inc(x_121); -x_122 = lean_ctor_get(x_120, 1); -lean_inc(x_122); -lean_dec(x_120); -x_123 = lean_ctor_get(x_121, 0); -lean_inc(x_123); -lean_dec(x_121); -x_124 = 1; -x_125 = lean_usize_add(x_5, x_124); -x_5 = x_125; -x_6 = x_123; -x_7 = x_122; +x_149 = l_Array_forInUnsafe_loop___at_Lean_finalizeImport___spec__8___lambda__1(x_10, x_2, x_144, x_130, x_147, x_148, x_7); +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, 0); +lean_inc(x_152); +lean_dec(x_150); +x_153 = 1; +x_154 = lean_usize_add(x_5, x_153); +x_5 = x_154; +x_6 = x_152; +x_7 = x_151; goto _start; } else { -lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; -lean_dec(x_115); -lean_dec(x_114); -x_127 = l_Lean_throwAlreadyImported___rarg(x_1, x_100, x_2, x_10, x_7); +lean_object* x_156; lean_object* x_157; uint8_t x_158; +x_156 = lean_ctor_get(x_145, 0); +lean_inc(x_156); +lean_dec(x_145); +x_157 = lean_ctor_get(x_146, 0); +lean_inc(x_157); +lean_dec(x_146); +x_158 = l___private_Lean_Environment_0__Lean_equivInfo(x_157, x_141); +lean_dec(x_141); +lean_dec(x_157); +if (x_158 == 0) +{ +lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; +lean_dec(x_156); +lean_dec(x_144); +x_159 = l_Lean_throwAlreadyImported___rarg(x_1, x_130, x_2, x_10, x_7); lean_dec(x_2); -x_128 = lean_ctor_get(x_127, 0); -lean_inc(x_128); -x_129 = lean_ctor_get(x_127, 1); -lean_inc(x_129); -if (lean_is_exclusive(x_127)) { - lean_ctor_release(x_127, 0); - lean_ctor_release(x_127, 1); - x_130 = x_127; +x_160 = lean_ctor_get(x_159, 0); +lean_inc(x_160); +x_161 = lean_ctor_get(x_159, 1); +lean_inc(x_161); +if (lean_is_exclusive(x_159)) { + lean_ctor_release(x_159, 0); + lean_ctor_release(x_159, 1); + x_162 = x_159; } else { - lean_dec_ref(x_127); - x_130 = lean_box(0); + lean_dec_ref(x_159); + x_162 = lean_box(0); } -if (lean_is_scalar(x_130)) { - x_131 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_162)) { + x_163 = lean_alloc_ctor(1, 2, 0); } else { - x_131 = x_130; + x_163 = x_162; +} +lean_ctor_set(x_163, 0, x_160); +lean_ctor_set(x_163, 1, x_161); +return x_163; +} +else +{ +lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; size_t x_169; size_t x_170; +x_164 = lean_box(0); +lean_inc(x_2); +x_165 = l_Array_forInUnsafe_loop___at_Lean_finalizeImport___spec__8___lambda__1(x_10, x_2, x_144, x_130, x_156, x_164, x_7); +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_ctor_get(x_166, 0); +lean_inc(x_168); +lean_dec(x_166); +x_169 = 1; +x_170 = lean_usize_add(x_5, x_169); +x_5 = x_170; +x_6 = x_168; +x_7 = x_167; +goto _start; } -lean_ctor_set(x_131, 0, x_128); -lean_ctor_set(x_131, 1, x_129); -return x_131; } } } @@ -13549,7 +13761,7 @@ x_7 = l_Lean_withImportModules___rarg(x_1, x_2, x_6, x_4, x_5); return x_7; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__2(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6351____spec__2(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { uint8_t x_5; @@ -13572,7 +13784,7 @@ return x_4; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__3(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_initFn____x40_Lean_Environment___hyg_6351____spec__3(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { uint8_t x_5; @@ -13610,7 +13822,7 @@ size_t x_15; size_t x_16; lean_object* x_17; x_15 = 0; x_16 = lean_usize_of_nat(x_7); lean_dec(x_7); -x_17 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__2(x_6, x_15, x_16, x_4); +x_17 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6351____spec__2(x_6, x_15, x_16, x_4); lean_dec(x_6); x_2 = x_11; x_4 = x_17; @@ -13624,7 +13836,7 @@ return x_4; } } } -LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6351____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; uint8_t x_5; @@ -13653,14 +13865,14 @@ size_t x_7; size_t x_8; lean_object* x_9; x_7 = 0; x_8 = lean_usize_of_nat(x_3); lean_dec(x_3); -x_9 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__3(x_2, x_7, x_8, x_1); +x_9 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6351____spec__3(x_2, x_7, x_8, x_1); lean_dec(x_2); return x_9; } } } } -LEAN_EXPORT lean_object* l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__4(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6351____spec__4(lean_object* x_1) { _start: { uint8_t x_2; @@ -13698,7 +13910,7 @@ return x_8; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__5(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_initFn____x40_Lean_Environment___hyg_6351____spec__5(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { uint8_t x_5; @@ -13724,7 +13936,7 @@ return x_4; } } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6214____lambda__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6351____lambda__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; @@ -13733,7 +13945,7 @@ x_4 = l_Lean_SMap_insert___at_Lean_NameSSet_insert___spec__1(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6214____lambda__2___closed__1() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6351____lambda__2___closed__1() { _start: { lean_object* x_1; lean_object* x_2; @@ -13742,7 +13954,7 @@ x_2 = l_Lean_mkHashMapImp___rarg(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6214____lambda__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6351____lambda__2(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; @@ -13753,15 +13965,15 @@ if (x_4 == 0) { 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_dec(x_2); -x_5 = l_Lean_initFn____x40_Lean_Environment___hyg_6214____lambda__2___closed__1; -x_6 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__1(x_5, x_1); +x_5 = l_Lean_initFn____x40_Lean_Environment___hyg_6351____lambda__2___closed__1; +x_6 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6351____spec__1(x_5, x_1); x_7 = 1; x_8 = l_Lean_mkEmptyEnvironment___lambda__1___closed__3; x_9 = lean_alloc_ctor(0, 2, 1); lean_ctor_set(x_9, 0, x_6); lean_ctor_set(x_9, 1, x_8); lean_ctor_set_uint8(x_9, sizeof(void*)*2, x_7); -x_10 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__4(x_9); +x_10 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6351____spec__4(x_9); return x_10; } else @@ -13772,15 +13984,15 @@ if (x_11 == 0) { lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_dec(x_2); -x_12 = l_Lean_initFn____x40_Lean_Environment___hyg_6214____lambda__2___closed__1; -x_13 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__1(x_12, x_1); +x_12 = l_Lean_initFn____x40_Lean_Environment___hyg_6351____lambda__2___closed__1; +x_13 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6351____spec__1(x_12, x_1); x_14 = 1; x_15 = l_Lean_mkEmptyEnvironment___lambda__1___closed__3; x_16 = lean_alloc_ctor(0, 2, 1); lean_ctor_set(x_16, 0, x_13); lean_ctor_set(x_16, 1, x_15); lean_ctor_set_uint8(x_16, sizeof(void*)*2, x_14); -x_17 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__4(x_16); +x_17 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6351____spec__4(x_16); return x_17; } else @@ -13789,23 +14001,23 @@ size_t x_18; size_t x_19; lean_object* x_20; lean_object* x_21; lean_object* x_2 x_18 = 0; x_19 = lean_usize_of_nat(x_2); lean_dec(x_2); -x_20 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__5(x_1, x_18, x_19, x_3); +x_20 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6351____spec__5(x_1, x_18, x_19, x_3); x_21 = l_Lean_mkHashMapImp___rarg(x_20); lean_dec(x_20); -x_22 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__1(x_21, x_1); +x_22 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6351____spec__1(x_21, x_1); x_23 = 1; x_24 = l_Lean_mkEmptyEnvironment___lambda__1___closed__3; x_25 = lean_alloc_ctor(0, 2, 1); lean_ctor_set(x_25, 0, x_22); lean_ctor_set(x_25, 1, x_24); lean_ctor_set_uint8(x_25, sizeof(void*)*2, x_23); -x_26 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__4(x_25); +x_26 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6351____spec__4(x_25); return x_26; } } } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__1() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6351____closed__1() { _start: { lean_object* x_1; @@ -13813,17 +14025,17 @@ x_1 = lean_mk_string_from_bytes("namespacesExt", 13); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__2() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6351____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___auto____x40_Lean_Environment___hyg_2443____closed__1; -x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_6351____closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__3() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6351____closed__3() { _start: { lean_object* x_1; @@ -13832,30 +14044,30 @@ lean_closure_set(x_1, 0, lean_box(0)); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__4() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6351____closed__4() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_initFn____x40_Lean_Environment___hyg_6214____lambda__1), 2, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_initFn____x40_Lean_Environment___hyg_6351____lambda__1), 2, 0); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__5() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6351____closed__5() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_initFn____x40_Lean_Environment___hyg_6214____lambda__2), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_initFn____x40_Lean_Environment___hyg_6351____lambda__2), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__6() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6351____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_initFn____x40_Lean_Environment___hyg_6214____closed__2; -x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__4; -x_3 = l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__5; -x_4 = l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__3; +x_1 = l_Lean_initFn____x40_Lean_Environment___hyg_6351____closed__2; +x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_6351____closed__4; +x_3 = l_Lean_initFn____x40_Lean_Environment___hyg_6351____closed__5; +x_4 = l_Lean_initFn____x40_Lean_Environment___hyg_6351____closed__3; x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -13864,16 +14076,16 @@ lean_ctor_set(x_5, 3, x_4); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6214_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6351_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__6; +x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_6351____closed__6; x_3 = l_Lean_registerSimplePersistentEnvExtension___rarg(x_2, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6351____spec__2___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; @@ -13881,12 +14093,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_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__2(x_1, x_5, x_6, x_4); +x_7 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6351____spec__2(x_1, x_5, x_6, x_4); lean_dec(x_1); return x_7; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__3___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_initFn____x40_Lean_Environment___hyg_6351____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; @@ -13894,12 +14106,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_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__3(x_1, x_5, x_6, x_4); +x_7 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6351____spec__3(x_1, x_5, x_6, x_4); lean_dec(x_1); return x_7; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__5___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_initFn____x40_Lean_Environment___hyg_6351____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; @@ -13907,7 +14119,7 @@ x_5 = lean_unbox_usize(x_2); lean_dec(x_2); x_6 = lean_unbox_usize(x_3); lean_dec(x_3); -x_7 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6214____spec__5(x_1, x_5, x_6, x_4); +x_7 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6351____spec__5(x_1, x_5, x_6, x_4); lean_dec(x_1); return x_7; } @@ -16081,21 +16293,21 @@ l_Lean_importModules___closed__1 = _init_l_Lean_importModules___closed__1(); lean_mark_persistent(l_Lean_importModules___closed__1); l_Lean_importModules___boxed__const__1 = _init_l_Lean_importModules___boxed__const__1(); lean_mark_persistent(l_Lean_importModules___boxed__const__1); -l_Lean_initFn____x40_Lean_Environment___hyg_6214____lambda__2___closed__1 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6214____lambda__2___closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6214____lambda__2___closed__1); -l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__1 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__1); -l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__2 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__2(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__2); -l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__3 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__3(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__3); -l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__4 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__4(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__4); -l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__5 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__5(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__5); -l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__6 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__6(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6214____closed__6); -if (builtin) {res = l_Lean_initFn____x40_Lean_Environment___hyg_6214_(lean_io_mk_world()); +l_Lean_initFn____x40_Lean_Environment___hyg_6351____lambda__2___closed__1 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6351____lambda__2___closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6351____lambda__2___closed__1); +l_Lean_initFn____x40_Lean_Environment___hyg_6351____closed__1 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6351____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6351____closed__1); +l_Lean_initFn____x40_Lean_Environment___hyg_6351____closed__2 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6351____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6351____closed__2); +l_Lean_initFn____x40_Lean_Environment___hyg_6351____closed__3 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6351____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6351____closed__3); +l_Lean_initFn____x40_Lean_Environment___hyg_6351____closed__4 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6351____closed__4(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6351____closed__4); +l_Lean_initFn____x40_Lean_Environment___hyg_6351____closed__5 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6351____closed__5(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6351____closed__5); +l_Lean_initFn____x40_Lean_Environment___hyg_6351____closed__6 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6351____closed__6(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6351____closed__6); +if (builtin) {res = l_Lean_initFn____x40_Lean_Environment___hyg_6351_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_namespacesExt = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_namespacesExt); diff --git a/stage0/stdlib/Lean/Linter/Builtin.c b/stage0/stdlib/Lean/Linter/Builtin.c index 927b7f451428..f6f41432617b 100644 --- a/stage0/stdlib/Lean/Linter/Builtin.c +++ b/stage0/stdlib/Lean/Linter/Builtin.c @@ -1239,7 +1239,7 @@ static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__ _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("def", 3); +x_1 = lean_mk_string_from_bytes("definition", 10); return x_1; } } diff --git a/stage0/stdlib/Lean/Linter/MissingDocs.c b/stage0/stdlib/Lean/Linter/MissingDocs.c index 3ca6af706499..27cb4e6bbd24 100644 --- a/stage0/stdlib/Lean/Linter/MissingDocs.c +++ b/stage0/stdlib/Lean/Linter/MissingDocs.c @@ -4083,7 +4083,7 @@ static lean_object* _init_l_Lean_Linter_MissingDocs_lintDeclHead___closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("def", 3); +x_1 = lean_mk_string_from_bytes("definition", 10); return x_1; } } diff --git a/stage0/stdlib/Lean/Meta.c b/stage0/stdlib/Lean/Meta.c index f29f846fddcb..00b3c2b64e27 100644 --- a/stage0/stdlib/Lean/Meta.c +++ b/stage0/stdlib/Lean/Meta.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Meta -// Imports: Lean.Meta.Basic Lean.Meta.LevelDefEq Lean.Meta.WHNF Lean.Meta.InferType Lean.Meta.FunInfo Lean.Meta.ExprDefEq Lean.Meta.DecLevel Lean.Meta.DiscrTree Lean.Meta.Reduce Lean.Meta.Instances Lean.Meta.AbstractMVars Lean.Meta.SynthInstance Lean.Meta.AppBuilder Lean.Meta.Tactic Lean.Meta.KAbstract Lean.Meta.RecursorInfo Lean.Meta.GeneralizeTelescope Lean.Meta.Match Lean.Meta.ReduceEval Lean.Meta.Closure Lean.Meta.AbstractNestedProofs Lean.Meta.ForEachExpr Lean.Meta.Transform Lean.Meta.PPGoal Lean.Meta.UnificationHint Lean.Meta.Inductive Lean.Meta.SizeOf Lean.Meta.IndPredBelow Lean.Meta.Coe Lean.Meta.CollectFVars Lean.Meta.GeneralizeVars Lean.Meta.Injective Lean.Meta.Structure Lean.Meta.Constructions Lean.Meta.CongrTheorems Lean.Meta.Eqns Lean.Meta.ExprLens Lean.Meta.ExprTraverse Lean.Meta.Eval Lean.Meta.CoeAttr Lean.Meta.Iterator Lean.Meta.LazyDiscrTree Lean.Meta.LitValues Lean.Meta.CheckTactic +// Imports: Lean.Meta.Basic Lean.Meta.LevelDefEq Lean.Meta.WHNF Lean.Meta.InferType Lean.Meta.FunInfo Lean.Meta.ExprDefEq Lean.Meta.DecLevel Lean.Meta.DiscrTree Lean.Meta.Reduce Lean.Meta.Instances Lean.Meta.AbstractMVars Lean.Meta.SynthInstance Lean.Meta.AppBuilder Lean.Meta.Tactic Lean.Meta.KAbstract Lean.Meta.RecursorInfo Lean.Meta.GeneralizeTelescope Lean.Meta.Match Lean.Meta.ReduceEval Lean.Meta.Closure Lean.Meta.AbstractNestedProofs Lean.Meta.ForEachExpr Lean.Meta.Transform Lean.Meta.PPGoal Lean.Meta.UnificationHint Lean.Meta.Inductive Lean.Meta.SizeOf Lean.Meta.IndPredBelow Lean.Meta.Coe Lean.Meta.CollectFVars Lean.Meta.GeneralizeVars Lean.Meta.Injective Lean.Meta.Structure Lean.Meta.Constructions Lean.Meta.CongrTheorems Lean.Meta.Eqns Lean.Meta.ExprLens Lean.Meta.ExprTraverse Lean.Meta.Eval Lean.Meta.CoeAttr Lean.Meta.Iterator Lean.Meta.LazyDiscrTree Lean.Meta.LitValues Lean.Meta.CheckTactic Lean.Meta.Canonicalizer #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -57,6 +57,7 @@ lean_object* initialize_Lean_Meta_Iterator(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_LazyDiscrTree(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_LitValues(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_CheckTactic(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Canonicalizer(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Meta(uint8_t builtin, lean_object* w) { lean_object * res; @@ -194,6 +195,9 @@ lean_dec_ref(res); res = initialize_Lean_Meta_CheckTactic(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_Meta_Canonicalizer(builtin, 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)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Meta/Canonicalizer.c b/stage0/stdlib/Lean/Meta/Canonicalizer.c new file mode 100644 index 000000000000..76cf38b808ed --- /dev/null +++ b/stage0/stdlib/Lean/Meta/Canonicalizer.c @@ -0,0 +1,5364 @@ +// Lean compiler output +// Module: Lean.Meta.Canonicalizer +// Imports: Lean.Util.ShareCommon Lean.Data.HashMap Lean.Meta.Basic Lean.Meta.FunInfo +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +LEAN_EXPORT lean_object* l_Lean_Meta_Canonicalizer_canon_unsafe__3(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_mk_empty_array_with_capacity(lean_object*); +lean_object* l_Lean_mkAppN(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Canonicalizer_State_cache___default___closed__1; +LEAN_EXPORT uint8_t l_Lean_AssocList_contains___at_Lean_Meta_Canonicalizer_canon_unsafe__2___spec__1(lean_object*, lean_object*); +lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_AssocList_replace___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__2___spec__6(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__2___spec__4(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*); +LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Meta_Canonicalizer_canon_unsafe__2___spec__4___at_Lean_Meta_Canonicalizer_canon_unsafe__2___spec__5(lean_object*, lean_object*); +lean_object* l_Lean_Expr_sort___override(lean_object*); +lean_object* lean_array_push(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_HashMapImp_expand___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__2___spec__2(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_Lean_AssocList_find_x3f___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__1___spec__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Canonicalizer_State_keys___default; +LEAN_EXPORT lean_object* l_Lean_Meta_Canonicalizer_CanonM_run___rarg(lean_object*, uint8_t, 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*); +lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_Meta_Canonicalizer_canon_unsafe__1___spec__2___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___lambda__2(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_Lean_Meta_Canonicalizer_instBEqExprVisited___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Canonicalizer_instInhabitedExprVisited; +lean_object* l_Lean_Level_ofNat(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Canonicalizer_canon___lambda__1(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_Lean_Meta_Canonicalizer_canon_unsafe__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_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__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_Canonicalizer_CanonM_run(lean_object*); +uint8_t l_Lean_Expr_hasMVar(lean_object*); +LEAN_EXPORT lean_object* l_Lean_HashMapImp_find_x3f___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__1___spec__1(lean_object*, lean_object*); +size_t lean_ptr_addr(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Canonicalizer_canon_unsafe__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___spec__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*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___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_Lean_AssocList_replace___at_Lean_Meta_Canonicalizer_canon_unsafe__2___spec__6(lean_object*, lean_object*, lean_object*); +lean_object* lean_st_ref_take(lean_object*, lean_object*); +uint8_t lean_expr_eqv(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Canonicalizer_instInhabitedState___closed__1; +lean_object* l_Lean_Meta_getFunInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint64_t lean_usize_to_uint64(size_t); +LEAN_EXPORT lean_object* l_Lean_Meta_Canonicalizer_State_keyToExprs___default; +LEAN_EXPORT lean_object* l_Lean_Meta_Canonicalizer_CanonM_run___rarg___boxed(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_Canonicalizer_canon_unsafe__1___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__1___spec__2___boxed(lean_object*, lean_object*); +uint8_t l_Lean_Expr_isMVar(lean_object*); +lean_object* l_Lean_Expr_forallE___override(lean_object*, lean_object*, lean_object*, uint8_t); +LEAN_EXPORT uint8_t l_Lean_Meta_Canonicalizer_instBEqExprVisited(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Canonicalizer_instInhabitedState; +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Meta_Canonicalizer_canon___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_st_ref_get(lean_object*, lean_object*); +lean_object* lean_st_mk_ref(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_AssocList_contains___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__2___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_shareCommon___rarg___boxed(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___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__2___spec__1___boxed(lean_object*, lean_object*); +extern lean_object* l_Lean_levelZero; +static lean_object* l_Lean_Meta_Canonicalizer_instInhabitedExprVisited___closed__1; +LEAN_EXPORT uint64_t l_Lean_Meta_Canonicalizer_instHashableExprVisited(lean_object*); +lean_object* l_ShareCommon_mkStateImpl(lean_object*); +lean_object* l_Lean_Expr_letE___override(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); +LEAN_EXPORT lean_object* l_Lean_HashMapImp_expand___at_Lean_Meta_Canonicalizer_canon_unsafe__2___spec__2(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Canonicalizer_canon___closed__1; +LEAN_EXPORT lean_object* l_Lean_HashMapImp_moveEntries___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__2___spec__3(lean_object*, lean_object*, lean_object*); +size_t lean_hashmap_mk_idx(lean_object*, uint64_t); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Canonicalizer_instHashableExprVisited___boxed(lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Meta_Canonicalizer_canon___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___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_shareCommon(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___lambda__2___closed__1; +static lean_object* l_Lean_Meta_Canonicalizer_State_keys___default___closed__1; +LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_Meta_Canonicalizer_canon_unsafe__1___spec__2(lean_object*, lean_object*); +static lean_object* l_Array_mapIdxM_map___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___spec__2___closed__2; +uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey(lean_object*, uint8_t, 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_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___spec__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_HashMapImp_moveEntries___at_Lean_Meta_Canonicalizer_canon_unsafe__2___spec__3(lean_object*, lean_object*, lean_object*); +lean_object* lean_nat_sub(lean_object*, lean_object*); +lean_object* l_Lean_Expr_getAppFn(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Canonicalizer_canon_unsafe__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_nat_mul(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Canonicalizer_State_cache___default; +lean_object* l_Lean_mkHashMapImp___rarg(lean_object*); +lean_object* l_Lean_instantiateMVarsCore(lean_object*, lean_object*); +static lean_object* l_Array_mapIdxM_map___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___spec__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Canonicalizer_canon_unsafe__2(lean_object*, lean_object*, lean_object*, uint8_t, 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___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_shareCommon___rarg(lean_object*, uint8_t, 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___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); +lean_object* lean_array_get_size(lean_object*); +extern lean_object* l_Lean_ShareCommon_objectFactory; +lean_object* lean_state_sharecommon(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Meta_Canonicalizer_canon_unsafe__2___spec__4(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___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*); +uint8_t lean_nat_dec_le(lean_object*, lean_object*); +lean_object* lean_nat_add(lean_object*, lean_object*); +lean_object* l_Lean_Expr_lam___override(lean_object*, lean_object*, lean_object*, uint8_t); +LEAN_EXPORT lean_object* l_Lean_Meta_Canonicalizer_canon___boxed(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_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__2___spec__4___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__2___spec__5(lean_object*, lean_object*); +uint8_t l_Lean_Meta_ParamInfo_isExplicit(lean_object*); +lean_object* lean_array_uset(lean_object*, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_AssocList_contains___at_Lean_Meta_Canonicalizer_canon_unsafe__2___spec__1___boxed(lean_object*, lean_object*); +lean_object* l___private_Lean_Data_HashMap_0__Lean_numBucketsForCapacity(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Canonicalizer_canon(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Canonicalizer_canon___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* _init_l_Lean_Meta_Canonicalizer_instInhabitedExprVisited___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = l_Lean_Expr_bvar___override(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Canonicalizer_instInhabitedExprVisited() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Meta_Canonicalizer_instInhabitedExprVisited___closed__1; +return x_1; +} +} +LEAN_EXPORT uint8_t l_Lean_Meta_Canonicalizer_instBEqExprVisited(lean_object* x_1, lean_object* x_2) { +_start: +{ +size_t x_3; size_t x_4; uint8_t x_5; +x_3 = lean_ptr_addr(x_1); +x_4 = lean_ptr_addr(x_2); +x_5 = lean_usize_dec_eq(x_3, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Canonicalizer_instBEqExprVisited___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_Meta_Canonicalizer_instBEqExprVisited(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT uint64_t l_Lean_Meta_Canonicalizer_instHashableExprVisited(lean_object* x_1) { +_start: +{ +size_t x_2; uint64_t x_3; +x_2 = lean_ptr_addr(x_1); +x_3 = lean_usize_to_uint64(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Canonicalizer_instHashableExprVisited___boxed(lean_object* x_1) { +_start: +{ +uint64_t x_2; lean_object* x_3; +x_2 = l_Lean_Meta_Canonicalizer_instHashableExprVisited(x_1); +lean_dec(x_1); +x_3 = lean_box_uint64(x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Canonicalizer_State_keys___default___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_ShareCommon_objectFactory; +x_2 = l_ShareCommon_mkStateImpl(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Canonicalizer_State_keys___default() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Meta_Canonicalizer_State_keys___default___closed__1; +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Canonicalizer_State_cache___default___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(8u); +x_2 = l_Lean_mkHashMapImp___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Canonicalizer_State_cache___default() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Meta_Canonicalizer_State_cache___default___closed__1; +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Canonicalizer_State_keyToExprs___default() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Meta_Canonicalizer_State_cache___default___closed__1; +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Canonicalizer_instInhabitedState___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Canonicalizer_State_keys___default___closed__1; +x_2 = l_Lean_Meta_Canonicalizer_State_cache___default___closed__1; +x_3 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +lean_ctor_set(x_3, 2, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Canonicalizer_instInhabitedState() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Meta_Canonicalizer_instInhabitedState___closed__1; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Canonicalizer_CanonM_run___rarg(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; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_8 = l_Lean_Meta_Canonicalizer_instInhabitedState___closed__1; +x_9 = lean_st_mk_ref(x_8, x_7); +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_box(x_2); +lean_inc(x_10); +x_13 = lean_apply_7(x_1, x_12, x_10, x_3, x_4, x_5, x_6, x_11); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = lean_st_ref_get(x_10, x_15); +lean_dec(x_10); +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_14); +return x_16; +} +else +{ +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_14); +lean_ctor_set(x_20, 1, x_19); +return x_20; +} +} +else +{ +uint8_t x_21; +lean_dec(x_10); +x_21 = !lean_is_exclusive(x_13); +if (x_21 == 0) +{ +return x_13; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_13, 0); +x_23 = lean_ctor_get(x_13, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_13); +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_Meta_Canonicalizer_CanonM_run(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_Canonicalizer_CanonM_run___rarg___boxed), 7, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Canonicalizer_CanonM_run___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: +{ +uint8_t x_8; lean_object* x_9; +x_8 = lean_unbox(x_2); +lean_dec(x_2); +x_9 = l_Lean_Meta_Canonicalizer_CanonM_run___rarg(x_1, x_8, x_3, x_4, x_5, x_6, x_7); +return x_9; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_shareCommon___rarg(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_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; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_13 = lean_ctor_get(x_10, 0); +x_14 = l_Lean_ShareCommon_objectFactory; +x_15 = lean_state_sharecommon(x_14, x_13, x_1); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +lean_ctor_set(x_10, 0, x_17); +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; +x_20 = lean_ctor_get(x_18, 0); +lean_dec(x_20); +lean_ctor_set(x_18, 0, x_16); +return x_18; +} +else +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_18, 1); +lean_inc(x_21); +lean_dec(x_18); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_16); +lean_ctor_set(x_22, 1, x_21); +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_ctor_get(x_10, 0); +x_24 = lean_ctor_get(x_10, 1); +x_25 = lean_ctor_get(x_10, 2); +lean_inc(x_25); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_10); +x_26 = l_Lean_ShareCommon_objectFactory; +x_27 = lean_state_sharecommon(x_26, x_23, x_1); +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_alloc_ctor(0, 3, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_24); +lean_ctor_set(x_30, 2, x_25); +x_31 = lean_st_ref_set(x_3, x_30, x_11); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +if (lean_is_exclusive(x_31)) { + lean_ctor_release(x_31, 0); + lean_ctor_release(x_31, 1); + x_33 = x_31; +} else { + lean_dec_ref(x_31); + x_33 = lean_box(0); +} +if (lean_is_scalar(x_33)) { + x_34 = lean_alloc_ctor(0, 2, 0); +} else { + x_34 = x_33; +} +lean_ctor_set(x_34, 0, x_28); +lean_ctor_set(x_34, 1, x_32); +return x_34; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_shareCommon(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_shareCommon___rarg___boxed), 8, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_shareCommon___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +uint8_t x_9; lean_object* x_10; +x_9 = lean_unbox(x_2); +lean_dec(x_2); +x_10 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_shareCommon___rarg(x_1, x_9, 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_10; +} +} +LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__1___spec__2(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; +x_3 = lean_box(0); +return x_3; +} +else +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; size_t x_7; size_t x_8; uint8_t x_9; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 1); +x_6 = lean_ctor_get(x_2, 2); +x_7 = lean_ptr_addr(x_4); +x_8 = lean_ptr_addr(x_1); +x_9 = lean_usize_dec_eq(x_7, x_8); +if (x_9 == 0) +{ +x_2 = x_6; +goto _start; +} +else +{ +lean_object* x_11; +lean_inc(x_5); +x_11 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_11, 0, x_5); +return x_11; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_HashMapImp_find_x3f___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__1___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; size_t x_5; uint64_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +lean_dec(x_1); +x_4 = lean_array_get_size(x_3); +x_5 = lean_ptr_addr(x_2); +x_6 = lean_usize_to_uint64(x_5); +x_7 = lean_hashmap_mk_idx(x_4, x_6); +x_8 = lean_array_uget(x_3, x_7); +lean_dec(x_3); +x_9 = l_Lean_AssocList_find_x3f___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__1___spec__2(x_2, x_8); +lean_dec(x_8); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; +x_3 = lean_ctor_get(x_2, 1); +lean_inc(x_3); +lean_dec(x_2); +x_4 = l_Lean_HashMapImp_find_x3f___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__1___spec__1(x_3, x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__1___spec__2___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_AssocList_find_x3f___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__1___spec__2(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT uint8_t l_Lean_AssocList_contains___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__2___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_3; +x_3 = 0; +return x_3; +} +else +{ +lean_object* x_4; lean_object* x_5; size_t x_6; size_t x_7; uint8_t x_8; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 2); +x_6 = lean_ptr_addr(x_4); +x_7 = lean_ptr_addr(x_1); +x_8 = lean_usize_dec_eq(x_6, x_7); +if (x_8 == 0) +{ +x_2 = x_5; +goto _start; +} +else +{ +uint8_t x_10; +x_10 = 1; +return x_10; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__2___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_3) == 0) +{ +lean_dec(x_1); +return x_2; +} +else +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint64_t x_9; size_t x_10; lean_object* x_11; lean_object* x_12; +x_5 = lean_ctor_get(x_3, 0); +x_6 = lean_ctor_get(x_3, 2); +x_7 = lean_array_get_size(x_2); +lean_inc(x_1); +lean_inc(x_5); +x_8 = lean_apply_1(x_1, x_5); +x_9 = lean_unbox_uint64(x_8); +lean_dec(x_8); +x_10 = lean_hashmap_mk_idx(x_7, x_9); +x_11 = lean_array_uget(x_2, x_10); +lean_ctor_set(x_3, 2, x_11); +x_12 = lean_array_uset(x_2, x_10, x_3); +x_2 = x_12; +x_3 = x_6; +goto _start; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint64_t x_19; size_t x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_14 = lean_ctor_get(x_3, 0); +x_15 = lean_ctor_get(x_3, 1); +x_16 = lean_ctor_get(x_3, 2); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_dec(x_3); +x_17 = lean_array_get_size(x_2); +lean_inc(x_1); +lean_inc(x_14); +x_18 = lean_apply_1(x_1, x_14); +x_19 = lean_unbox_uint64(x_18); +lean_dec(x_18); +x_20 = lean_hashmap_mk_idx(x_17, x_19); +x_21 = lean_array_uget(x_2, x_20); +x_22 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_22, 0, x_14); +lean_ctor_set(x_22, 1, x_15); +lean_ctor_set(x_22, 2, x_21); +x_23 = lean_array_uset(x_2, x_20, x_22); +x_2 = x_23; +x_3 = x_16; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__2___spec__4___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__2___spec__5(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +return x_1; +} +else +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_2); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; size_t x_7; uint64_t x_8; size_t x_9; lean_object* x_10; lean_object* x_11; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 2); +x_6 = lean_array_get_size(x_1); +x_7 = lean_ptr_addr(x_4); +x_8 = lean_usize_to_uint64(x_7); +x_9 = lean_hashmap_mk_idx(x_6, x_8); +x_10 = lean_array_uget(x_1, x_9); +lean_ctor_set(x_2, 2, x_10); +x_11 = lean_array_uset(x_1, x_9, x_2); +x_1 = x_11; +x_2 = x_5; +goto _start; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; size_t x_17; uint64_t x_18; size_t x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_13 = lean_ctor_get(x_2, 0); +x_14 = lean_ctor_get(x_2, 1); +x_15 = lean_ctor_get(x_2, 2); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_dec(x_2); +x_16 = lean_array_get_size(x_1); +x_17 = lean_ptr_addr(x_13); +x_18 = lean_usize_to_uint64(x_17); +x_19 = lean_hashmap_mk_idx(x_16, x_18); +x_20 = lean_array_uget(x_1, x_19); +x_21 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_21, 0, x_13); +lean_ctor_set(x_21, 1, x_14); +lean_ctor_set(x_21, 2, x_20); +x_22 = lean_array_uset(x_1, x_19, x_21); +x_1 = x_22; +x_2 = x_15; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_HashMapImp_moveEntries___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__2___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_array_get_size(x_2); +x_5 = lean_nat_dec_lt(x_1, x_4); +lean_dec(x_4); +if (x_5 == 0) +{ +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_6 = lean_array_fget(x_2, x_1); +x_7 = lean_box(0); +x_8 = lean_array_fset(x_2, x_1, x_7); +x_9 = l_Lean_AssocList_foldlM___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__2___spec__4___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__2___spec__5(x_3, x_6); +x_10 = lean_unsigned_to_nat(1u); +x_11 = lean_nat_add(x_1, x_10); +lean_dec(x_1); +x_1 = x_11; +x_2 = x_8; +x_3 = x_9; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Lean_HashMapImp_expand___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__2___spec__2(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_3 = lean_array_get_size(x_2); +x_4 = lean_unsigned_to_nat(2u); +x_5 = lean_nat_mul(x_3, x_4); +lean_dec(x_3); +x_6 = lean_box(0); +x_7 = lean_mk_array(x_5, x_6); +x_8 = lean_unsigned_to_nat(0u); +x_9 = l_Lean_HashMapImp_moveEntries___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__2___spec__3(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___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__2___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_4; +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(0); +return x_4; +} +else +{ +uint8_t x_5; +x_5 = !lean_is_exclusive(x_3); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; uint8_t x_11; +x_6 = lean_ctor_get(x_3, 0); +x_7 = lean_ctor_get(x_3, 1); +x_8 = lean_ctor_get(x_3, 2); +x_9 = lean_ptr_addr(x_6); +x_10 = lean_ptr_addr(x_1); +x_11 = lean_usize_dec_eq(x_9, x_10); +if (x_11 == 0) +{ +lean_object* x_12; +x_12 = l_Lean_AssocList_replace___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__2___spec__6(x_1, x_2, x_8); +lean_ctor_set(x_3, 2, x_12); +return x_3; +} +else +{ +lean_dec(x_7); +lean_dec(x_6); +lean_ctor_set(x_3, 1, x_2); +lean_ctor_set(x_3, 0, x_1); +return x_3; +} +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; size_t x_16; size_t x_17; uint8_t x_18; +x_13 = lean_ctor_get(x_3, 0); +x_14 = lean_ctor_get(x_3, 1); +x_15 = lean_ctor_get(x_3, 2); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_dec(x_3); +x_16 = lean_ptr_addr(x_13); +x_17 = lean_ptr_addr(x_1); +x_18 = lean_usize_dec_eq(x_16, x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +x_19 = l_Lean_AssocList_replace___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__2___spec__6(x_1, x_2, x_15); +x_20 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_20, 0, x_13); +lean_ctor_set(x_20, 1, x_14); +lean_ctor_set(x_20, 2, x_19); +return x_20; +} +else +{ +lean_object* x_21; +lean_dec(x_14); +lean_dec(x_13); +x_21 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_21, 0, x_1); +lean_ctor_set(x_21, 1, x_2); +lean_ctor_set(x_21, 2, x_15); +return x_21; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__2(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_10 = lean_st_ref_take(x_4, x_9); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_11, 1); +lean_inc(x_12); +x_13 = lean_ctor_get(x_10, 1); +lean_inc(x_13); +lean_dec(x_10); +x_14 = !lean_is_exclusive(x_11); +if (x_14 == 0) +{ +lean_object* x_15; uint8_t x_16; +x_15 = lean_ctor_get(x_11, 1); +lean_dec(x_15); +x_16 = !lean_is_exclusive(x_12); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; size_t x_20; uint64_t x_21; size_t x_22; lean_object* x_23; uint8_t x_24; +x_17 = lean_ctor_get(x_12, 0); +x_18 = lean_ctor_get(x_12, 1); +x_19 = lean_array_get_size(x_18); +x_20 = lean_ptr_addr(x_1); +x_21 = lean_usize_to_uint64(x_20); +lean_inc(x_19); +x_22 = lean_hashmap_mk_idx(x_19, x_21); +x_23 = lean_array_uget(x_18, x_22); +x_24 = l_Lean_AssocList_contains___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__2___spec__1(x_1, x_23); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; +x_25 = lean_unsigned_to_nat(1u); +x_26 = lean_nat_add(x_17, x_25); +lean_dec(x_17); +x_27 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_27, 0, x_1); +lean_ctor_set(x_27, 1, x_2); +lean_ctor_set(x_27, 2, x_23); +x_28 = lean_array_uset(x_18, x_22, x_27); +x_29 = l___private_Lean_Data_HashMap_0__Lean_numBucketsForCapacity(x_26); +x_30 = lean_nat_dec_le(x_29, x_19); +lean_dec(x_19); +lean_dec(x_29); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; uint8_t x_33; +lean_free_object(x_12); +x_31 = l_Lean_HashMapImp_expand___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__2___spec__2(x_26, x_28); +lean_ctor_set(x_11, 1, x_31); +x_32 = lean_st_ref_set(x_4, x_11, x_13); +x_33 = !lean_is_exclusive(x_32); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; +x_34 = lean_ctor_get(x_32, 0); +lean_dec(x_34); +x_35 = lean_box(0); +lean_ctor_set(x_32, 0, x_35); +return x_32; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_32, 1); +lean_inc(x_36); +lean_dec(x_32); +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_36); +return x_38; +} +} +else +{ +lean_object* x_39; uint8_t x_40; +lean_ctor_set(x_12, 1, x_28); +lean_ctor_set(x_12, 0, x_26); +x_39 = lean_st_ref_set(x_4, x_11, x_13); +x_40 = !lean_is_exclusive(x_39); +if (x_40 == 0) +{ +lean_object* x_41; lean_object* x_42; +x_41 = lean_ctor_get(x_39, 0); +lean_dec(x_41); +x_42 = lean_box(0); +lean_ctor_set(x_39, 0, x_42); +return x_39; +} +else +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_43 = lean_ctor_get(x_39, 1); +lean_inc(x_43); +lean_dec(x_39); +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; +} +} +} +else +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; +lean_dec(x_19); +x_46 = l_Lean_AssocList_replace___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__2___spec__6(x_1, x_2, x_23); +x_47 = lean_array_uset(x_18, x_22, x_46); +lean_ctor_set(x_12, 1, x_47); +x_48 = lean_st_ref_set(x_4, x_11, x_13); +x_49 = !lean_is_exclusive(x_48); +if (x_49 == 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; +} +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; +} +} +} +else +{ +lean_object* x_55; lean_object* x_56; lean_object* x_57; size_t x_58; uint64_t x_59; size_t x_60; lean_object* x_61; uint8_t x_62; +x_55 = lean_ctor_get(x_12, 0); +x_56 = lean_ctor_get(x_12, 1); +lean_inc(x_56); +lean_inc(x_55); +lean_dec(x_12); +x_57 = lean_array_get_size(x_56); +x_58 = lean_ptr_addr(x_1); +x_59 = lean_usize_to_uint64(x_58); +lean_inc(x_57); +x_60 = lean_hashmap_mk_idx(x_57, x_59); +x_61 = lean_array_uget(x_56, x_60); +x_62 = l_Lean_AssocList_contains___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__2___spec__1(x_1, x_61); +if (x_62 == 0) +{ +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_63 = lean_unsigned_to_nat(1u); +x_64 = lean_nat_add(x_55, x_63); +lean_dec(x_55); +x_65 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_65, 0, x_1); +lean_ctor_set(x_65, 1, x_2); +lean_ctor_set(x_65, 2, x_61); +x_66 = lean_array_uset(x_56, x_60, x_65); +x_67 = l___private_Lean_Data_HashMap_0__Lean_numBucketsForCapacity(x_64); +x_68 = lean_nat_dec_le(x_67, x_57); +lean_dec(x_57); +lean_dec(x_67); +if (x_68 == 0) +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_69 = l_Lean_HashMapImp_expand___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__2___spec__2(x_64, x_66); +lean_ctor_set(x_11, 1, x_69); +x_70 = lean_st_ref_set(x_4, x_11, x_13); +x_71 = lean_ctor_get(x_70, 1); +lean_inc(x_71); +if (lean_is_exclusive(x_70)) { + lean_ctor_release(x_70, 0); + lean_ctor_release(x_70, 1); + x_72 = x_70; +} else { + lean_dec_ref(x_70); + x_72 = lean_box(0); +} +x_73 = lean_box(0); +if (lean_is_scalar(x_72)) { + x_74 = lean_alloc_ctor(0, 2, 0); +} else { + x_74 = x_72; +} +lean_ctor_set(x_74, 0, x_73); +lean_ctor_set(x_74, 1, x_71); +return x_74; +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; +x_75 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_75, 0, x_64); +lean_ctor_set(x_75, 1, x_66); +lean_ctor_set(x_11, 1, x_75); +x_76 = lean_st_ref_set(x_4, x_11, x_13); +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); +} +x_79 = lean_box(0); +if (lean_is_scalar(x_78)) { + x_80 = lean_alloc_ctor(0, 2, 0); +} else { + x_80 = x_78; +} +lean_ctor_set(x_80, 0, x_79); +lean_ctor_set(x_80, 1, x_77); +return x_80; +} +} +else +{ +lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; +lean_dec(x_57); +x_81 = l_Lean_AssocList_replace___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__2___spec__6(x_1, x_2, x_61); +x_82 = lean_array_uset(x_56, x_60, x_81); +x_83 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_83, 0, x_55); +lean_ctor_set(x_83, 1, x_82); +lean_ctor_set(x_11, 1, x_83); +x_84 = lean_st_ref_set(x_4, x_11, x_13); +x_85 = lean_ctor_get(x_84, 1); +lean_inc(x_85); +if (lean_is_exclusive(x_84)) { + lean_ctor_release(x_84, 0); + lean_ctor_release(x_84, 1); + x_86 = x_84; +} else { + lean_dec_ref(x_84); + x_86 = lean_box(0); +} +x_87 = lean_box(0); +if (lean_is_scalar(x_86)) { + x_88 = lean_alloc_ctor(0, 2, 0); +} else { + x_88 = x_86; +} +lean_ctor_set(x_88, 0, x_87); +lean_ctor_set(x_88, 1, x_85); +return x_88; +} +} +} +else +{ +lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; size_t x_95; uint64_t x_96; size_t x_97; lean_object* x_98; uint8_t x_99; +x_89 = lean_ctor_get(x_11, 0); +x_90 = lean_ctor_get(x_11, 2); +lean_inc(x_90); +lean_inc(x_89); +lean_dec(x_11); +x_91 = lean_ctor_get(x_12, 0); +lean_inc(x_91); +x_92 = lean_ctor_get(x_12, 1); +lean_inc(x_92); +if (lean_is_exclusive(x_12)) { + lean_ctor_release(x_12, 0); + lean_ctor_release(x_12, 1); + x_93 = x_12; +} else { + lean_dec_ref(x_12); + x_93 = lean_box(0); +} +x_94 = lean_array_get_size(x_92); +x_95 = lean_ptr_addr(x_1); +x_96 = lean_usize_to_uint64(x_95); +lean_inc(x_94); +x_97 = lean_hashmap_mk_idx(x_94, x_96); +x_98 = lean_array_uget(x_92, x_97); +x_99 = l_Lean_AssocList_contains___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__2___spec__1(x_1, x_98); +if (x_99 == 0) +{ +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; uint8_t x_105; +x_100 = lean_unsigned_to_nat(1u); +x_101 = lean_nat_add(x_91, x_100); +lean_dec(x_91); +x_102 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_102, 0, x_1); +lean_ctor_set(x_102, 1, x_2); +lean_ctor_set(x_102, 2, x_98); +x_103 = lean_array_uset(x_92, x_97, x_102); +x_104 = l___private_Lean_Data_HashMap_0__Lean_numBucketsForCapacity(x_101); +x_105 = lean_nat_dec_le(x_104, x_94); +lean_dec(x_94); +lean_dec(x_104); +if (x_105 == 0) +{ +lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +lean_dec(x_93); +x_106 = l_Lean_HashMapImp_expand___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__2___spec__2(x_101, x_103); +x_107 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_107, 0, x_89); +lean_ctor_set(x_107, 1, x_106); +lean_ctor_set(x_107, 2, x_90); +x_108 = lean_st_ref_set(x_4, x_107, x_13); +x_109 = lean_ctor_get(x_108, 1); +lean_inc(x_109); +if (lean_is_exclusive(x_108)) { + lean_ctor_release(x_108, 0); + lean_ctor_release(x_108, 1); + x_110 = x_108; +} else { + lean_dec_ref(x_108); + x_110 = lean_box(0); +} +x_111 = lean_box(0); +if (lean_is_scalar(x_110)) { + x_112 = lean_alloc_ctor(0, 2, 0); +} else { + x_112 = x_110; +} +lean_ctor_set(x_112, 0, x_111); +lean_ctor_set(x_112, 1, x_109); +return x_112; +} +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; +if (lean_is_scalar(x_93)) { + x_113 = lean_alloc_ctor(0, 2, 0); +} else { + x_113 = x_93; +} +lean_ctor_set(x_113, 0, x_101); +lean_ctor_set(x_113, 1, x_103); +x_114 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_114, 0, x_89); +lean_ctor_set(x_114, 1, x_113); +lean_ctor_set(x_114, 2, x_90); +x_115 = lean_st_ref_set(x_4, x_114, x_13); +x_116 = lean_ctor_get(x_115, 1); +lean_inc(x_116); +if (lean_is_exclusive(x_115)) { + lean_ctor_release(x_115, 0); + lean_ctor_release(x_115, 1); + x_117 = x_115; +} else { + lean_dec_ref(x_115); + x_117 = lean_box(0); +} +x_118 = lean_box(0); +if (lean_is_scalar(x_117)) { + x_119 = lean_alloc_ctor(0, 2, 0); +} else { + x_119 = x_117; +} +lean_ctor_set(x_119, 0, x_118); +lean_ctor_set(x_119, 1, x_116); +return x_119; +} +} +else +{ +lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; +lean_dec(x_94); +x_120 = l_Lean_AssocList_replace___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__2___spec__6(x_1, x_2, x_98); +x_121 = lean_array_uset(x_92, x_97, x_120); +if (lean_is_scalar(x_93)) { + x_122 = lean_alloc_ctor(0, 2, 0); +} else { + x_122 = x_93; +} +lean_ctor_set(x_122, 0, x_91); +lean_ctor_set(x_122, 1, x_121); +x_123 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_123, 0, x_89); +lean_ctor_set(x_123, 1, x_122); +lean_ctor_set(x_123, 2, x_90); +x_124 = lean_st_ref_set(x_4, x_123, x_13); +x_125 = lean_ctor_get(x_124, 1); +lean_inc(x_125); +if (lean_is_exclusive(x_124)) { + lean_ctor_release(x_124, 0); + lean_ctor_release(x_124, 1); + x_126 = x_124; +} else { + lean_dec_ref(x_124); + x_126 = lean_box(0); +} +x_127 = lean_box(0); +if (lean_is_scalar(x_126)) { + x_128 = lean_alloc_ctor(0, 2, 0); +} else { + x_128 = x_126; +} +lean_ctor_set(x_128, 0, x_127); +lean_ctor_set(x_128, 1, x_125); +return x_128; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_AssocList_contains___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__2___spec__1___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___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__2___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_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__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_3); +lean_dec(x_3); +x_11 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__2(x_1, x_2, x_10, 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_11; +} +} +LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___spec__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) { +_start: +{ +uint8_t x_9; +x_9 = l_Lean_Expr_hasMVar(x_1); +if (x_9 == 0) +{ +lean_object* x_10; +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_1); +lean_ctor_set(x_10, 1, x_8); +return x_10; +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_11 = lean_st_ref_get(x_5, 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 = lean_ctor_get(x_12, 0); +lean_inc(x_14); +lean_dec(x_12); +x_15 = l_Lean_instantiateMVarsCore(x_14, x_1); +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_st_ref_take(x_5, x_13); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = !lean_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_5, x_19, x_20); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) +{ +lean_object* x_25; +x_25 = lean_ctor_get(x_23, 0); +lean_dec(x_25); +lean_ctor_set(x_23, 0, x_16); +return x_23; +} +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_16); +lean_ctor_set(x_27, 1, x_26); +return x_27; +} +} +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_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_5, 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); +} +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_16); +lean_ctor_set(x_35, 1, x_33); +return x_35; +} +} +} +} +static lean_object* _init_l_Array_mapIdxM_map___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___spec__2___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = l_Lean_Level_ofNat(x_1); +return x_2; +} +} +static lean_object* _init_l_Array_mapIdxM_map___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___spec__2___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_mapIdxM_map___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___spec__2___closed__1; +x_2 = l_Lean_Expr_sort___override(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___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, 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) { +_start: +{ +lean_object* x_15; uint8_t x_16; +x_15 = lean_unsigned_to_nat(0u); +x_16 = lean_nat_dec_eq(x_4, 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; +x_17 = lean_unsigned_to_nat(1u); +x_18 = lean_nat_sub(x_4, x_17); +lean_dec(x_4); +x_19 = lean_array_fget(x_3, x_5); +x_20 = lean_ctor_get(x_2, 0); +x_21 = lean_array_get_size(x_20); +x_22 = lean_nat_dec_lt(x_5, x_21); +lean_dec(x_21); +if (x_22 == 0) +{ +lean_object* x_23; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +x_23 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey(x_19, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +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_nat_add(x_5, x_17); +lean_dec(x_5); +x_27 = lean_array_push(x_7, x_24); +x_4 = x_18; +x_5 = x_26; +x_6 = lean_box(0); +x_7 = x_27; +x_14 = x_25; +goto _start; +} +else +{ +uint8_t x_29; +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_7); +lean_dec(x_5); +x_29 = !lean_is_exclusive(x_23); +if (x_29 == 0) +{ +return x_23; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_23, 0); +x_31 = lean_ctor_get(x_23, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_23); +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; +} +} +} +else +{ +lean_object* x_33; uint8_t x_34; +x_33 = lean_array_fget(x_20, x_5); +x_34 = l_Lean_Meta_ParamInfo_isExplicit(x_33); +lean_dec(x_33); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +lean_dec(x_19); +x_35 = lean_nat_add(x_5, x_17); +lean_dec(x_5); +x_36 = l_Array_mapIdxM_map___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___spec__2___closed__2; +x_37 = lean_array_push(x_7, x_36); +x_4 = x_18; +x_5 = x_35; +x_6 = lean_box(0); +x_7 = x_37; +goto _start; +} +else +{ +lean_object* x_39; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +x_39 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey(x_19, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_39) == 0) +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +x_41 = lean_ctor_get(x_39, 1); +lean_inc(x_41); +lean_dec(x_39); +x_42 = lean_nat_add(x_5, x_17); +lean_dec(x_5); +x_43 = lean_array_push(x_7, x_40); +x_4 = x_18; +x_5 = x_42; +x_6 = lean_box(0); +x_7 = x_43; +x_14 = x_41; +goto _start; +} +else +{ +uint8_t x_45; +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_7); +lean_dec(x_5); +x_45 = !lean_is_exclusive(x_39); +if (x_45 == 0) +{ +return x_39; +} +else +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_ctor_get(x_39, 0); +x_47 = lean_ctor_get(x_39, 1); +lean_inc(x_47); +lean_inc(x_46); +lean_dec(x_39); +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_49; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_4); +x_49 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_49, 0, x_7); +lean_ctor_set(x_49, 1, x_14); +return x_49; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___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) { +_start: +{ +lean_object* x_10; uint8_t x_11; +lean_inc(x_2); +x_10 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) +{ +lean_object* x_12; +x_12 = lean_ctor_get(x_10, 0); +lean_dec(x_12); +lean_ctor_set(x_10, 0, x_2); +return x_10; +} +else +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_ctor_get(x_10, 1); +lean_inc(x_13); +lean_dec(x_10); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_2); +lean_ctor_set(x_14, 1, x_13); +return x_14; +} +} +} +static lean_object* _init_l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___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); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___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, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_1); +x_12 = l_Lean_Meta_getFunInfo(x_1, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_12) == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +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_unsigned_to_nat(0u); +x_16 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_2, x_15); +x_17 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___lambda__2___closed__1; +lean_inc(x_16); +x_18 = lean_mk_array(x_16, x_17); +x_19 = lean_unsigned_to_nat(1u); +x_20 = lean_nat_sub(x_16, x_19); +lean_dec(x_16); +lean_inc(x_2); +x_21 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_2, x_18, x_20); +x_22 = lean_array_get_size(x_21); +x_23 = lean_mk_empty_array_with_capacity(x_22); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_24 = l_Array_mapIdxM_map___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___spec__2(x_2, x_13, x_21, x_22, x_15, lean_box(0), x_23, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +lean_dec(x_21); +lean_dec(x_13); +lean_dec(x_2); +if (lean_obj_tag(x_24) == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +lean_dec(x_24); +x_27 = l_Lean_mkAppN(x_1, x_25); +x_28 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_shareCommon___rarg(x_27, x_5, x_6, 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); +lean_inc(x_30); +lean_dec(x_28); +x_31 = lean_box(x_5); +x_32 = lean_apply_8(x_3, x_29, x_31, x_6, x_7, x_8, x_9, x_10, x_30); +return x_32; +} +else +{ +uint8_t x_33; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_1); +x_33 = !lean_is_exclusive(x_24); +if (x_33 == 0) +{ +return x_24; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_24, 0); +x_35 = lean_ctor_get(x_24, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_24); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; +} +} +} +else +{ +uint8_t x_37; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_2); +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___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey(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_9; uint8_t x_10; +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; +x_11 = lean_ctor_get(x_9, 0); +x_12 = lean_ctor_get(x_9, 1); +lean_inc(x_1); +x_13 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__1(x_1, x_11); +if (lean_obj_tag(x_13) == 0) +{ +lean_free_object(x_9); +switch (lean_obj_tag(x_1)) { +case 2: +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +lean_inc(x_1); +x_14 = l_Lean_instantiateMVars___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_12); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = lean_expr_eqv(x_15, x_1); +if (x_17 == 0) +{ +lean_object* x_18; +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_18 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey(x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_16); +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_18, 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_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___lambda__1(x_1, x_19, 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); +return x_21; +} +else +{ +uint8_t x_22; +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_18); +if (x_22 == 0) +{ +return x_18; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_18, 0); +x_24 = lean_ctor_get(x_18, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_18); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; +} +} +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +lean_dec(x_15); +lean_inc(x_1); +x_26 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_shareCommon___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_16); +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +lean_dec(x_26); +x_29 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___lambda__1(x_1, x_27, x_2, x_3, x_4, x_5, x_6, x_7, x_28); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_29; +} +} +case 5: +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +lean_inc(x_1); +x_30 = lean_alloc_closure((void*)(l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___lambda__1___boxed), 9, 1); +lean_closure_set(x_30, 0, x_1); +x_31 = l_Lean_Expr_getAppFn(x_1); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_32 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey(x_31, x_2, x_3, x_4, x_5, x_6, x_7, x_12); +if (lean_obj_tag(x_32) == 0) +{ +lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +x_35 = l_Lean_Expr_isMVar(x_33); +if (x_35 == 0) +{ +lean_object* x_36; lean_object* x_37; +x_36 = lean_box(0); +x_37 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___lambda__2(x_33, x_1, x_30, x_36, x_2, x_3, x_4, x_5, x_6, x_7, x_34); +return x_37; +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +lean_inc(x_1); +x_38 = l_Lean_instantiateMVars___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_34); +x_39 = lean_ctor_get(x_38, 0); +lean_inc(x_39); +x_40 = lean_ctor_get(x_38, 1); +lean_inc(x_40); +lean_dec(x_38); +x_41 = lean_expr_eqv(x_39, x_1); +if (x_41 == 0) +{ +lean_object* x_42; +lean_dec(x_33); +lean_dec(x_30); +lean_dec(x_1); +x_42 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey(x_39, x_2, x_3, x_4, x_5, x_6, x_7, x_40); +if (lean_obj_tag(x_42) == 0) +{ +uint8_t x_43; +x_43 = !lean_is_exclusive(x_42); +if (x_43 == 0) +{ +return x_42; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_42, 0); +x_45 = lean_ctor_get(x_42, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; +} +} +else +{ +uint8_t x_47; +x_47 = !lean_is_exclusive(x_42); +if (x_47 == 0) +{ +return x_42; +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_42, 0); +x_49 = lean_ctor_get(x_42, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_42); +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 +{ +lean_object* x_51; lean_object* x_52; +lean_dec(x_39); +x_51 = lean_box(0); +x_52 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___lambda__2(x_33, x_1, x_30, x_51, x_2, x_3, x_4, x_5, x_6, x_7, x_40); +return x_52; +} +} +} +else +{ +uint8_t x_53; +lean_dec(x_30); +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_53 = !lean_is_exclusive(x_32); +if (x_53 == 0) +{ +return x_32; +} +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_54 = lean_ctor_get(x_32, 0); +x_55 = lean_ctor_get(x_32, 1); +lean_inc(x_55); +lean_inc(x_54); +lean_dec(x_32); +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; +} +} +} +case 6: +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; uint8_t x_60; lean_object* x_61; +x_57 = lean_ctor_get(x_1, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_1, 1); +lean_inc(x_58); +x_59 = lean_ctor_get(x_1, 2); +lean_inc(x_59); +x_60 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_61 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey(x_58, x_2, x_3, x_4, x_5, x_6, x_7, x_12); +if (lean_obj_tag(x_61) == 0) +{ +lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_62 = lean_ctor_get(x_61, 0); +lean_inc(x_62); +x_63 = lean_ctor_get(x_61, 1); +lean_inc(x_63); +lean_dec(x_61); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_64 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey(x_59, x_2, x_3, x_4, x_5, x_6, x_7, 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; 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_Expr_lam___override(x_57, x_62, x_65, x_60); +x_68 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_shareCommon___rarg(x_67, x_2, x_3, x_4, x_5, x_6, x_7, x_66); +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___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___lambda__1(x_1, x_69, x_2, x_3, x_4, x_5, x_6, x_7, x_70); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_71; +} +else +{ +uint8_t x_72; +lean_dec(x_62); +lean_dec(x_57); +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_72 = !lean_is_exclusive(x_64); +if (x_72 == 0) +{ +return x_64; +} +else +{ +lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_73 = lean_ctor_get(x_64, 0); +x_74 = lean_ctor_get(x_64, 1); +lean_inc(x_74); +lean_inc(x_73); +lean_dec(x_64); +x_75 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_75, 0, x_73); +lean_ctor_set(x_75, 1, x_74); +return x_75; +} +} +} +else +{ +uint8_t x_76; +lean_dec(x_59); +lean_dec(x_57); +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_76 = !lean_is_exclusive(x_61); +if (x_76 == 0) +{ +return x_61; +} +else +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_77 = lean_ctor_get(x_61, 0); +x_78 = lean_ctor_get(x_61, 1); +lean_inc(x_78); +lean_inc(x_77); +lean_dec(x_61); +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; +} +} +} +case 7: +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; uint8_t x_83; lean_object* x_84; +x_80 = lean_ctor_get(x_1, 0); +lean_inc(x_80); +x_81 = lean_ctor_get(x_1, 1); +lean_inc(x_81); +x_82 = lean_ctor_get(x_1, 2); +lean_inc(x_82); +x_83 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_84 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey(x_81, x_2, x_3, x_4, x_5, x_6, x_7, x_12); +if (lean_obj_tag(x_84) == 0) +{ +lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_85 = lean_ctor_get(x_84, 0); +lean_inc(x_85); +x_86 = lean_ctor_get(x_84, 1); +lean_inc(x_86); +lean_dec(x_84); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_87 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey(x_82, x_2, x_3, x_4, x_5, x_6, x_7, x_86); +if (lean_obj_tag(x_87) == 0) +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); +x_90 = l_Lean_Expr_forallE___override(x_80, x_85, x_88, x_83); +x_91 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_shareCommon___rarg(x_90, x_2, x_3, x_4, x_5, x_6, x_7, x_89); +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___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___lambda__1(x_1, x_92, x_2, x_3, x_4, x_5, x_6, x_7, x_93); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_94; +} +else +{ +uint8_t x_95; +lean_dec(x_85); +lean_dec(x_80); +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_95 = !lean_is_exclusive(x_87); +if (x_95 == 0) +{ +return x_87; +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_87, 0); +x_97 = lean_ctor_get(x_87, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_87); +x_98 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +return x_98; +} +} +} +else +{ +uint8_t x_99; +lean_dec(x_82); +lean_dec(x_80); +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_99 = !lean_is_exclusive(x_84); +if (x_99 == 0) +{ +return x_84; +} +else +{ +lean_object* x_100; lean_object* x_101; lean_object* x_102; +x_100 = lean_ctor_get(x_84, 0); +x_101 = lean_ctor_get(x_84, 1); +lean_inc(x_101); +lean_inc(x_100); +lean_dec(x_84); +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; +} +} +} +case 8: +{ +lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; uint8_t x_107; lean_object* x_108; +x_103 = lean_ctor_get(x_1, 0); +lean_inc(x_103); +x_104 = lean_ctor_get(x_1, 1); +lean_inc(x_104); +x_105 = lean_ctor_get(x_1, 2); +lean_inc(x_105); +x_106 = lean_ctor_get(x_1, 3); +lean_inc(x_106); +x_107 = lean_ctor_get_uint8(x_1, sizeof(void*)*4 + 8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_108 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey(x_104, x_2, x_3, x_4, x_5, x_6, x_7, x_12); +if (lean_obj_tag(x_108) == 0) +{ +lean_object* x_109; lean_object* x_110; lean_object* x_111; +x_109 = lean_ctor_get(x_108, 0); +lean_inc(x_109); +x_110 = lean_ctor_get(x_108, 1); +lean_inc(x_110); +lean_dec(x_108); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_111 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey(x_105, x_2, x_3, x_4, x_5, x_6, x_7, x_110); +if (lean_obj_tag(x_111) == 0) +{ +lean_object* x_112; lean_object* x_113; lean_object* x_114; +x_112 = lean_ctor_get(x_111, 0); +lean_inc(x_112); +x_113 = lean_ctor_get(x_111, 1); +lean_inc(x_113); +lean_dec(x_111); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_114 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey(x_106, x_2, x_3, x_4, x_5, x_6, x_7, x_113); +if (lean_obj_tag(x_114) == 0) +{ +lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; +x_115 = lean_ctor_get(x_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_Expr_letE___override(x_103, x_109, x_112, x_115, x_107); +x_118 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_shareCommon___rarg(x_117, x_2, x_3, x_4, x_5, x_6, x_7, x_116); +x_119 = lean_ctor_get(x_118, 0); +lean_inc(x_119); +x_120 = lean_ctor_get(x_118, 1); +lean_inc(x_120); +lean_dec(x_118); +x_121 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___lambda__1(x_1, x_119, x_2, x_3, x_4, x_5, x_6, x_7, x_120); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_121; +} +else +{ +uint8_t x_122; +lean_dec(x_112); +lean_dec(x_109); +lean_dec(x_103); +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_122 = !lean_is_exclusive(x_114); +if (x_122 == 0) +{ +return x_114; +} +else +{ +lean_object* x_123; lean_object* x_124; lean_object* x_125; +x_123 = lean_ctor_get(x_114, 0); +x_124 = lean_ctor_get(x_114, 1); +lean_inc(x_124); +lean_inc(x_123); +lean_dec(x_114); +x_125 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_125, 0, x_123); +lean_ctor_set(x_125, 1, x_124); +return x_125; +} +} +} +else +{ +uint8_t x_126; +lean_dec(x_109); +lean_dec(x_106); +lean_dec(x_103); +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_126 = !lean_is_exclusive(x_111); +if (x_126 == 0) +{ +return x_111; +} +else +{ +lean_object* x_127; lean_object* x_128; lean_object* x_129; +x_127 = lean_ctor_get(x_111, 0); +x_128 = lean_ctor_get(x_111, 1); +lean_inc(x_128); +lean_inc(x_127); +lean_dec(x_111); +x_129 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_129, 0, x_127); +lean_ctor_set(x_129, 1, x_128); +return x_129; +} +} +} +else +{ +uint8_t x_130; +lean_dec(x_106); +lean_dec(x_105); +lean_dec(x_103); +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_130 = !lean_is_exclusive(x_108); +if (x_130 == 0) +{ +return x_108; +} +else +{ +lean_object* x_131; lean_object* x_132; lean_object* x_133; +x_131 = lean_ctor_get(x_108, 0); +x_132 = lean_ctor_get(x_108, 1); +lean_inc(x_132); +lean_inc(x_131); +lean_dec(x_108); +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 10: +{ +lean_object* x_134; lean_object* x_135; +x_134 = lean_ctor_get(x_1, 1); +lean_inc(x_134); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_135 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey(x_134, x_2, x_3, x_4, x_5, x_6, x_7, x_12); +if (lean_obj_tag(x_135) == 0) +{ +lean_object* x_136; lean_object* x_137; lean_object* x_138; +x_136 = lean_ctor_get(x_135, 0); +lean_inc(x_136); +x_137 = lean_ctor_get(x_135, 1); +lean_inc(x_137); +lean_dec(x_135); +x_138 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___lambda__1(x_1, x_136, x_2, x_3, x_4, x_5, x_6, x_7, x_137); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_138; +} +else +{ +uint8_t x_139; +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_139 = !lean_is_exclusive(x_135); +if (x_139 == 0) +{ +return x_135; +} +else +{ +lean_object* x_140; lean_object* x_141; lean_object* x_142; +x_140 = lean_ctor_get(x_135, 0); +x_141 = lean_ctor_get(x_135, 1); +lean_inc(x_141); +lean_inc(x_140); +lean_dec(x_135); +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; +} +} +} +case 11: +{ +lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; +x_143 = lean_ctor_get(x_1, 0); +lean_inc(x_143); +x_144 = lean_ctor_get(x_1, 1); +lean_inc(x_144); +x_145 = lean_ctor_get(x_1, 2); +lean_inc(x_145); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_146 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey(x_145, x_2, x_3, x_4, x_5, x_6, x_7, x_12); +if (lean_obj_tag(x_146) == 0) +{ +lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; +x_147 = lean_ctor_get(x_146, 0); +lean_inc(x_147); +x_148 = lean_ctor_get(x_146, 1); +lean_inc(x_148); +lean_dec(x_146); +x_149 = l_Lean_Expr_proj___override(x_143, x_144, x_147); +x_150 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_shareCommon___rarg(x_149, x_2, x_3, x_4, x_5, x_6, x_7, x_148); +x_151 = lean_ctor_get(x_150, 0); +lean_inc(x_151); +x_152 = lean_ctor_get(x_150, 1); +lean_inc(x_152); +lean_dec(x_150); +x_153 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___lambda__1(x_1, x_151, x_2, x_3, x_4, x_5, x_6, x_7, x_152); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_153; +} +else +{ +uint8_t x_154; +lean_dec(x_144); +lean_dec(x_143); +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_154 = !lean_is_exclusive(x_146); +if (x_154 == 0) +{ +return x_146; +} +else +{ +lean_object* x_155; lean_object* x_156; lean_object* x_157; +x_155 = lean_ctor_get(x_146, 0); +x_156 = lean_ctor_get(x_146, 1); +lean_inc(x_156); +lean_inc(x_155); +lean_dec(x_146); +x_157 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_157, 0, x_155); +lean_ctor_set(x_157, 1, x_156); +return x_157; +} +} +} +default: +{ +lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; +lean_inc(x_1); +x_158 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_shareCommon___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_12); +x_159 = lean_ctor_get(x_158, 0); +lean_inc(x_159); +x_160 = lean_ctor_get(x_158, 1); +lean_inc(x_160); +lean_dec(x_158); +x_161 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___lambda__1(x_1, x_159, x_2, x_3, x_4, x_5, x_6, x_7, x_160); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_161; +} +} +} +else +{ +lean_object* x_162; +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_162 = lean_ctor_get(x_13, 0); +lean_inc(x_162); +lean_dec(x_13); +lean_ctor_set(x_9, 0, x_162); +return x_9; +} +} +else +{ +lean_object* x_163; lean_object* x_164; lean_object* x_165; +x_163 = lean_ctor_get(x_9, 0); +x_164 = lean_ctor_get(x_9, 1); +lean_inc(x_164); +lean_inc(x_163); +lean_dec(x_9); +lean_inc(x_1); +x_165 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey_unsafe__1(x_1, x_163); +if (lean_obj_tag(x_165) == 0) +{ +switch (lean_obj_tag(x_1)) { +case 2: +{ +lean_object* x_166; lean_object* x_167; lean_object* x_168; uint8_t x_169; +lean_inc(x_1); +x_166 = l_Lean_instantiateMVars___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_164); +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_expr_eqv(x_167, x_1); +if (x_169 == 0) +{ +lean_object* x_170; +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_170 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey(x_167, x_2, x_3, x_4, x_5, x_6, x_7, x_168); +if (lean_obj_tag(x_170) == 0) +{ +lean_object* x_171; lean_object* x_172; lean_object* x_173; +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 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___lambda__1(x_1, x_171, x_2, x_3, x_4, x_5, x_6, x_7, x_172); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_173; +} +else +{ +lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; +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_174 = lean_ctor_get(x_170, 0); +lean_inc(x_174); +x_175 = lean_ctor_get(x_170, 1); +lean_inc(x_175); +if (lean_is_exclusive(x_170)) { + lean_ctor_release(x_170, 0); + lean_ctor_release(x_170, 1); + x_176 = x_170; +} else { + lean_dec_ref(x_170); + x_176 = lean_box(0); +} +if (lean_is_scalar(x_176)) { + x_177 = lean_alloc_ctor(1, 2, 0); +} else { + x_177 = x_176; +} +lean_ctor_set(x_177, 0, x_174); +lean_ctor_set(x_177, 1, x_175); +return x_177; +} +} +else +{ +lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; +lean_dec(x_167); +lean_inc(x_1); +x_178 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_shareCommon___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_168); +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 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___lambda__1(x_1, x_179, x_2, x_3, x_4, x_5, x_6, x_7, x_180); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_181; +} +} +case 5: +{ +lean_object* x_182; lean_object* x_183; lean_object* x_184; +lean_inc(x_1); +x_182 = lean_alloc_closure((void*)(l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___lambda__1___boxed), 9, 1); +lean_closure_set(x_182, 0, x_1); +x_183 = l_Lean_Expr_getAppFn(x_1); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_184 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey(x_183, x_2, x_3, x_4, x_5, x_6, x_7, x_164); +if (lean_obj_tag(x_184) == 0) +{ +lean_object* x_185; lean_object* x_186; uint8_t x_187; +x_185 = lean_ctor_get(x_184, 0); +lean_inc(x_185); +x_186 = lean_ctor_get(x_184, 1); +lean_inc(x_186); +lean_dec(x_184); +x_187 = l_Lean_Expr_isMVar(x_185); +if (x_187 == 0) +{ +lean_object* x_188; lean_object* x_189; +x_188 = lean_box(0); +x_189 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___lambda__2(x_185, x_1, x_182, x_188, x_2, x_3, x_4, x_5, x_6, x_7, x_186); +return x_189; +} +else +{ +lean_object* x_190; lean_object* x_191; lean_object* x_192; uint8_t x_193; +lean_inc(x_1); +x_190 = l_Lean_instantiateMVars___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_186); +x_191 = lean_ctor_get(x_190, 0); +lean_inc(x_191); +x_192 = lean_ctor_get(x_190, 1); +lean_inc(x_192); +lean_dec(x_190); +x_193 = lean_expr_eqv(x_191, x_1); +if (x_193 == 0) +{ +lean_object* x_194; +lean_dec(x_185); +lean_dec(x_182); +lean_dec(x_1); +x_194 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey(x_191, x_2, x_3, x_4, x_5, x_6, x_7, x_192); +if (lean_obj_tag(x_194) == 0) +{ +lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; +x_195 = lean_ctor_get(x_194, 0); +lean_inc(x_195); +x_196 = lean_ctor_get(x_194, 1); +lean_inc(x_196); +if (lean_is_exclusive(x_194)) { + lean_ctor_release(x_194, 0); + lean_ctor_release(x_194, 1); + x_197 = x_194; +} else { + lean_dec_ref(x_194); + x_197 = lean_box(0); +} +if (lean_is_scalar(x_197)) { + x_198 = lean_alloc_ctor(0, 2, 0); +} else { + x_198 = x_197; +} +lean_ctor_set(x_198, 0, x_195); +lean_ctor_set(x_198, 1, x_196); +return x_198; +} +else +{ +lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; +x_199 = lean_ctor_get(x_194, 0); +lean_inc(x_199); +x_200 = lean_ctor_get(x_194, 1); +lean_inc(x_200); +if (lean_is_exclusive(x_194)) { + lean_ctor_release(x_194, 0); + lean_ctor_release(x_194, 1); + x_201 = x_194; +} else { + lean_dec_ref(x_194); + x_201 = lean_box(0); +} +if (lean_is_scalar(x_201)) { + x_202 = lean_alloc_ctor(1, 2, 0); +} else { + x_202 = x_201; +} +lean_ctor_set(x_202, 0, x_199); +lean_ctor_set(x_202, 1, x_200); +return x_202; +} +} +else +{ +lean_object* x_203; lean_object* x_204; +lean_dec(x_191); +x_203 = lean_box(0); +x_204 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___lambda__2(x_185, x_1, x_182, x_203, x_2, x_3, x_4, x_5, x_6, x_7, x_192); +return x_204; +} +} +} +else +{ +lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; +lean_dec(x_182); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_205 = lean_ctor_get(x_184, 0); +lean_inc(x_205); +x_206 = lean_ctor_get(x_184, 1); +lean_inc(x_206); +if (lean_is_exclusive(x_184)) { + lean_ctor_release(x_184, 0); + lean_ctor_release(x_184, 1); + x_207 = x_184; +} else { + lean_dec_ref(x_184); + x_207 = lean_box(0); +} +if (lean_is_scalar(x_207)) { + x_208 = lean_alloc_ctor(1, 2, 0); +} else { + x_208 = x_207; +} +lean_ctor_set(x_208, 0, x_205); +lean_ctor_set(x_208, 1, x_206); +return x_208; +} +} +case 6: +{ +lean_object* x_209; lean_object* x_210; lean_object* x_211; uint8_t x_212; lean_object* x_213; +x_209 = lean_ctor_get(x_1, 0); +lean_inc(x_209); +x_210 = lean_ctor_get(x_1, 1); +lean_inc(x_210); +x_211 = lean_ctor_get(x_1, 2); +lean_inc(x_211); +x_212 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_213 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey(x_210, x_2, x_3, x_4, x_5, x_6, x_7, x_164); +if (lean_obj_tag(x_213) == 0) +{ +lean_object* x_214; lean_object* x_215; lean_object* x_216; +x_214 = lean_ctor_get(x_213, 0); +lean_inc(x_214); +x_215 = lean_ctor_get(x_213, 1); +lean_inc(x_215); +lean_dec(x_213); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_216 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey(x_211, x_2, x_3, x_4, x_5, x_6, x_7, x_215); +if (lean_obj_tag(x_216) == 0) +{ +lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; +x_217 = lean_ctor_get(x_216, 0); +lean_inc(x_217); +x_218 = lean_ctor_get(x_216, 1); +lean_inc(x_218); +lean_dec(x_216); +x_219 = l_Lean_Expr_lam___override(x_209, x_214, x_217, x_212); +x_220 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_shareCommon___rarg(x_219, x_2, x_3, x_4, x_5, x_6, x_7, x_218); +x_221 = lean_ctor_get(x_220, 0); +lean_inc(x_221); +x_222 = lean_ctor_get(x_220, 1); +lean_inc(x_222); +lean_dec(x_220); +x_223 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___lambda__1(x_1, x_221, x_2, x_3, x_4, x_5, x_6, x_7, x_222); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_223; +} +else +{ +lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; +lean_dec(x_214); +lean_dec(x_209); +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_224 = lean_ctor_get(x_216, 0); +lean_inc(x_224); +x_225 = lean_ctor_get(x_216, 1); +lean_inc(x_225); +if (lean_is_exclusive(x_216)) { + lean_ctor_release(x_216, 0); + lean_ctor_release(x_216, 1); + x_226 = x_216; +} else { + lean_dec_ref(x_216); + x_226 = lean_box(0); +} +if (lean_is_scalar(x_226)) { + x_227 = lean_alloc_ctor(1, 2, 0); +} else { + x_227 = x_226; +} +lean_ctor_set(x_227, 0, x_224); +lean_ctor_set(x_227, 1, x_225); +return x_227; +} +} +else +{ +lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; +lean_dec(x_211); +lean_dec(x_209); +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_228 = lean_ctor_get(x_213, 0); +lean_inc(x_228); +x_229 = lean_ctor_get(x_213, 1); +lean_inc(x_229); +if (lean_is_exclusive(x_213)) { + lean_ctor_release(x_213, 0); + lean_ctor_release(x_213, 1); + x_230 = x_213; +} else { + lean_dec_ref(x_213); + x_230 = lean_box(0); +} +if (lean_is_scalar(x_230)) { + x_231 = lean_alloc_ctor(1, 2, 0); +} else { + x_231 = x_230; +} +lean_ctor_set(x_231, 0, x_228); +lean_ctor_set(x_231, 1, x_229); +return x_231; +} +} +case 7: +{ +lean_object* x_232; lean_object* x_233; lean_object* x_234; uint8_t x_235; lean_object* x_236; +x_232 = lean_ctor_get(x_1, 0); +lean_inc(x_232); +x_233 = lean_ctor_get(x_1, 1); +lean_inc(x_233); +x_234 = lean_ctor_get(x_1, 2); +lean_inc(x_234); +x_235 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_236 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey(x_233, x_2, x_3, x_4, x_5, x_6, x_7, x_164); +if (lean_obj_tag(x_236) == 0) +{ +lean_object* x_237; lean_object* x_238; lean_object* x_239; +x_237 = lean_ctor_get(x_236, 0); +lean_inc(x_237); +x_238 = lean_ctor_get(x_236, 1); +lean_inc(x_238); +lean_dec(x_236); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_239 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey(x_234, x_2, x_3, x_4, x_5, x_6, x_7, x_238); +if (lean_obj_tag(x_239) == 0) +{ +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; +x_240 = lean_ctor_get(x_239, 0); +lean_inc(x_240); +x_241 = lean_ctor_get(x_239, 1); +lean_inc(x_241); +lean_dec(x_239); +x_242 = l_Lean_Expr_forallE___override(x_232, x_237, x_240, x_235); +x_243 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_shareCommon___rarg(x_242, x_2, x_3, x_4, x_5, x_6, x_7, x_241); +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_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___lambda__1(x_1, x_244, x_2, x_3, x_4, x_5, x_6, x_7, x_245); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_246; +} +else +{ +lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; +lean_dec(x_237); +lean_dec(x_232); +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_247 = lean_ctor_get(x_239, 0); +lean_inc(x_247); +x_248 = lean_ctor_get(x_239, 1); +lean_inc(x_248); +if (lean_is_exclusive(x_239)) { + lean_ctor_release(x_239, 0); + lean_ctor_release(x_239, 1); + x_249 = x_239; +} else { + lean_dec_ref(x_239); + 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; +} +} +else +{ +lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; +lean_dec(x_234); +lean_dec(x_232); +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_251 = lean_ctor_get(x_236, 0); +lean_inc(x_251); +x_252 = lean_ctor_get(x_236, 1); +lean_inc(x_252); +if (lean_is_exclusive(x_236)) { + lean_ctor_release(x_236, 0); + lean_ctor_release(x_236, 1); + x_253 = x_236; +} else { + lean_dec_ref(x_236); + x_253 = lean_box(0); +} +if (lean_is_scalar(x_253)) { + x_254 = lean_alloc_ctor(1, 2, 0); +} else { + x_254 = x_253; +} +lean_ctor_set(x_254, 0, x_251); +lean_ctor_set(x_254, 1, x_252); +return x_254; +} +} +case 8: +{ +lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; uint8_t x_259; lean_object* x_260; +x_255 = lean_ctor_get(x_1, 0); +lean_inc(x_255); +x_256 = lean_ctor_get(x_1, 1); +lean_inc(x_256); +x_257 = lean_ctor_get(x_1, 2); +lean_inc(x_257); +x_258 = lean_ctor_get(x_1, 3); +lean_inc(x_258); +x_259 = lean_ctor_get_uint8(x_1, sizeof(void*)*4 + 8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_260 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey(x_256, x_2, x_3, x_4, x_5, x_6, x_7, x_164); +if (lean_obj_tag(x_260) == 0) +{ +lean_object* x_261; lean_object* x_262; lean_object* x_263; +x_261 = lean_ctor_get(x_260, 0); +lean_inc(x_261); +x_262 = lean_ctor_get(x_260, 1); +lean_inc(x_262); +lean_dec(x_260); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_263 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey(x_257, x_2, x_3, x_4, x_5, x_6, x_7, x_262); +if (lean_obj_tag(x_263) == 0) +{ +lean_object* x_264; lean_object* x_265; lean_object* x_266; +x_264 = lean_ctor_get(x_263, 0); +lean_inc(x_264); +x_265 = lean_ctor_get(x_263, 1); +lean_inc(x_265); +lean_dec(x_263); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_266 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey(x_258, x_2, x_3, x_4, x_5, x_6, x_7, x_265); +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); +x_268 = lean_ctor_get(x_266, 1); +lean_inc(x_268); +lean_dec(x_266); +x_269 = l_Lean_Expr_letE___override(x_255, x_261, x_264, x_267, x_259); +x_270 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_shareCommon___rarg(x_269, x_2, x_3, x_4, x_5, x_6, x_7, x_268); +x_271 = lean_ctor_get(x_270, 0); +lean_inc(x_271); +x_272 = lean_ctor_get(x_270, 1); +lean_inc(x_272); +lean_dec(x_270); +x_273 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___lambda__1(x_1, x_271, x_2, x_3, x_4, x_5, x_6, x_7, x_272); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_273; +} +else +{ +lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; +lean_dec(x_264); +lean_dec(x_261); +lean_dec(x_255); +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_274 = lean_ctor_get(x_266, 0); +lean_inc(x_274); +x_275 = lean_ctor_get(x_266, 1); +lean_inc(x_275); +if (lean_is_exclusive(x_266)) { + lean_ctor_release(x_266, 0); + lean_ctor_release(x_266, 1); + x_276 = x_266; +} else { + lean_dec_ref(x_266); + x_276 = lean_box(0); +} +if (lean_is_scalar(x_276)) { + x_277 = lean_alloc_ctor(1, 2, 0); +} else { + x_277 = x_276; +} +lean_ctor_set(x_277, 0, x_274); +lean_ctor_set(x_277, 1, x_275); +return x_277; +} +} +else +{ +lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; +lean_dec(x_261); +lean_dec(x_258); +lean_dec(x_255); +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_278 = lean_ctor_get(x_263, 0); +lean_inc(x_278); +x_279 = lean_ctor_get(x_263, 1); +lean_inc(x_279); +if (lean_is_exclusive(x_263)) { + lean_ctor_release(x_263, 0); + lean_ctor_release(x_263, 1); + x_280 = x_263; +} else { + lean_dec_ref(x_263); + x_280 = lean_box(0); +} +if (lean_is_scalar(x_280)) { + x_281 = lean_alloc_ctor(1, 2, 0); +} else { + x_281 = x_280; +} +lean_ctor_set(x_281, 0, x_278); +lean_ctor_set(x_281, 1, x_279); +return x_281; +} +} +else +{ +lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; +lean_dec(x_258); +lean_dec(x_257); +lean_dec(x_255); +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_282 = lean_ctor_get(x_260, 0); +lean_inc(x_282); +x_283 = lean_ctor_get(x_260, 1); +lean_inc(x_283); +if (lean_is_exclusive(x_260)) { + lean_ctor_release(x_260, 0); + lean_ctor_release(x_260, 1); + x_284 = x_260; +} else { + lean_dec_ref(x_260); + x_284 = lean_box(0); +} +if (lean_is_scalar(x_284)) { + x_285 = lean_alloc_ctor(1, 2, 0); +} else { + x_285 = x_284; +} +lean_ctor_set(x_285, 0, x_282); +lean_ctor_set(x_285, 1, x_283); +return x_285; +} +} +case 10: +{ +lean_object* x_286; lean_object* x_287; +x_286 = lean_ctor_get(x_1, 1); +lean_inc(x_286); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_287 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey(x_286, x_2, x_3, x_4, x_5, x_6, x_7, x_164); +if (lean_obj_tag(x_287) == 0) +{ +lean_object* x_288; lean_object* x_289; lean_object* x_290; +x_288 = lean_ctor_get(x_287, 0); +lean_inc(x_288); +x_289 = lean_ctor_get(x_287, 1); +lean_inc(x_289); +lean_dec(x_287); +x_290 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___lambda__1(x_1, x_288, x_2, x_3, x_4, x_5, x_6, x_7, x_289); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_290; +} +else +{ +lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; +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_291 = lean_ctor_get(x_287, 0); +lean_inc(x_291); +x_292 = lean_ctor_get(x_287, 1); +lean_inc(x_292); +if (lean_is_exclusive(x_287)) { + lean_ctor_release(x_287, 0); + lean_ctor_release(x_287, 1); + x_293 = x_287; +} else { + lean_dec_ref(x_287); + x_293 = lean_box(0); +} +if (lean_is_scalar(x_293)) { + x_294 = lean_alloc_ctor(1, 2, 0); +} else { + x_294 = x_293; +} +lean_ctor_set(x_294, 0, x_291); +lean_ctor_set(x_294, 1, x_292); +return x_294; +} +} +case 11: +{ +lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; +x_295 = lean_ctor_get(x_1, 0); +lean_inc(x_295); +x_296 = lean_ctor_get(x_1, 1); +lean_inc(x_296); +x_297 = lean_ctor_get(x_1, 2); +lean_inc(x_297); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_298 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey(x_297, x_2, x_3, x_4, x_5, x_6, x_7, x_164); +if (lean_obj_tag(x_298) == 0) +{ +lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; +x_299 = lean_ctor_get(x_298, 0); +lean_inc(x_299); +x_300 = lean_ctor_get(x_298, 1); +lean_inc(x_300); +lean_dec(x_298); +x_301 = l_Lean_Expr_proj___override(x_295, x_296, x_299); +x_302 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_shareCommon___rarg(x_301, x_2, x_3, x_4, x_5, x_6, x_7, x_300); +x_303 = lean_ctor_get(x_302, 0); +lean_inc(x_303); +x_304 = lean_ctor_get(x_302, 1); +lean_inc(x_304); +lean_dec(x_302); +x_305 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___lambda__1(x_1, x_303, x_2, x_3, x_4, x_5, x_6, x_7, x_304); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_305; +} +else +{ +lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; +lean_dec(x_296); +lean_dec(x_295); +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_306 = lean_ctor_get(x_298, 0); +lean_inc(x_306); +x_307 = lean_ctor_get(x_298, 1); +lean_inc(x_307); +if (lean_is_exclusive(x_298)) { + lean_ctor_release(x_298, 0); + lean_ctor_release(x_298, 1); + x_308 = x_298; +} else { + lean_dec_ref(x_298); + x_308 = lean_box(0); +} +if (lean_is_scalar(x_308)) { + x_309 = lean_alloc_ctor(1, 2, 0); +} else { + x_309 = x_308; +} +lean_ctor_set(x_309, 0, x_306); +lean_ctor_set(x_309, 1, x_307); +return x_309; +} +} +default: +{ +lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; +lean_inc(x_1); +x_310 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_shareCommon___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_164); +x_311 = lean_ctor_get(x_310, 0); +lean_inc(x_311); +x_312 = lean_ctor_get(x_310, 1); +lean_inc(x_312); +lean_dec(x_310); +x_313 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___lambda__1(x_1, x_311, x_2, x_3, x_4, x_5, x_6, x_7, x_312); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_313; +} +} +} +else +{ +lean_object* x_314; lean_object* x_315; +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_314 = lean_ctor_get(x_165, 0); +lean_inc(x_314); +lean_dec(x_165); +x_315 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_315, 0, x_314); +lean_ctor_set(x_315, 1, x_164); +return x_315; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___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: +{ +uint8_t x_9; lean_object* x_10; +x_9 = lean_unbox(x_2); +lean_dec(x_2); +x_10 = l_Lean_instantiateMVars___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___spec__1(x_1, x_9, 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_10; +} +} +LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +uint8_t x_15; lean_object* x_16; +x_15 = lean_unbox(x_8); +lean_dec(x_8); +x_16 = l_Array_mapIdxM_map___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_15, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_16; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___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_10; lean_object* x_11; +x_10 = lean_unbox(x_3); +lean_dec(x_3); +x_11 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___lambda__1(x_1, x_2, x_10, 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_11; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +uint8_t x_12; lean_object* x_13; +x_12 = lean_unbox(x_5); +lean_dec(x_5); +x_13 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___lambda__2(x_1, x_2, x_3, x_4, x_12, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_4); +return x_13; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___boxed(lean_object* x_1, lean_object* x_2, 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___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey(x_1, x_9, x_3, x_4, x_5, x_6, x_7, x_8); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_Meta_Canonicalizer_canon_unsafe__1___spec__2(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; +x_3 = lean_box(0); +return x_3; +} +else +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; size_t x_7; size_t x_8; uint8_t x_9; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 1); +x_6 = lean_ctor_get(x_2, 2); +x_7 = lean_ptr_addr(x_4); +x_8 = lean_ptr_addr(x_1); +x_9 = lean_usize_dec_eq(x_7, x_8); +if (x_9 == 0) +{ +x_2 = x_6; +goto _start; +} +else +{ +lean_object* x_11; +lean_inc(x_5); +x_11 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_11, 0, x_5); +return x_11; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_HashMapImp_find_x3f___at_Lean_Meta_Canonicalizer_canon_unsafe__1___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; size_t x_5; uint64_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +lean_dec(x_1); +x_4 = lean_array_get_size(x_3); +x_5 = lean_ptr_addr(x_2); +x_6 = lean_usize_to_uint64(x_5); +x_7 = lean_hashmap_mk_idx(x_4, x_6); +x_8 = lean_array_uget(x_3, x_7); +lean_dec(x_3); +x_9 = l_Lean_AssocList_find_x3f___at_Lean_Meta_Canonicalizer_canon_unsafe__1___spec__2(x_2, x_8); +lean_dec(x_8); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Canonicalizer_canon_unsafe__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; +x_3 = lean_ctor_get(x_2, 2); +lean_inc(x_3); +lean_dec(x_2); +x_4 = l_Lean_HashMapImp_find_x3f___at_Lean_Meta_Canonicalizer_canon_unsafe__1___spec__1(x_3, x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_Meta_Canonicalizer_canon_unsafe__1___spec__2___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_AssocList_find_x3f___at_Lean_Meta_Canonicalizer_canon_unsafe__1___spec__2(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT uint8_t l_Lean_AssocList_contains___at_Lean_Meta_Canonicalizer_canon_unsafe__2___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_3; +x_3 = 0; +return x_3; +} +else +{ +lean_object* x_4; lean_object* x_5; size_t x_6; size_t x_7; uint8_t x_8; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 2); +x_6 = lean_ptr_addr(x_4); +x_7 = lean_ptr_addr(x_1); +x_8 = lean_usize_dec_eq(x_6, x_7); +if (x_8 == 0) +{ +x_2 = x_5; +goto _start; +} +else +{ +uint8_t x_10; +x_10 = 1; +return x_10; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Meta_Canonicalizer_canon_unsafe__2___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_3) == 0) +{ +lean_dec(x_1); +return x_2; +} +else +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint64_t x_9; size_t x_10; lean_object* x_11; lean_object* x_12; +x_5 = lean_ctor_get(x_3, 0); +x_6 = lean_ctor_get(x_3, 2); +x_7 = lean_array_get_size(x_2); +lean_inc(x_1); +lean_inc(x_5); +x_8 = lean_apply_1(x_1, x_5); +x_9 = lean_unbox_uint64(x_8); +lean_dec(x_8); +x_10 = lean_hashmap_mk_idx(x_7, x_9); +x_11 = lean_array_uget(x_2, x_10); +lean_ctor_set(x_3, 2, x_11); +x_12 = lean_array_uset(x_2, x_10, x_3); +x_2 = x_12; +x_3 = x_6; +goto _start; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint64_t x_19; size_t x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_14 = lean_ctor_get(x_3, 0); +x_15 = lean_ctor_get(x_3, 1); +x_16 = lean_ctor_get(x_3, 2); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_dec(x_3); +x_17 = lean_array_get_size(x_2); +lean_inc(x_1); +lean_inc(x_14); +x_18 = lean_apply_1(x_1, x_14); +x_19 = lean_unbox_uint64(x_18); +lean_dec(x_18); +x_20 = lean_hashmap_mk_idx(x_17, x_19); +x_21 = lean_array_uget(x_2, x_20); +x_22 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_22, 0, x_14); +lean_ctor_set(x_22, 1, x_15); +lean_ctor_set(x_22, 2, x_21); +x_23 = lean_array_uset(x_2, x_20, x_22); +x_2 = x_23; +x_3 = x_16; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Meta_Canonicalizer_canon_unsafe__2___spec__4___at_Lean_Meta_Canonicalizer_canon_unsafe__2___spec__5(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +return x_1; +} +else +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_2); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; size_t x_7; uint64_t x_8; size_t x_9; lean_object* x_10; lean_object* x_11; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 2); +x_6 = lean_array_get_size(x_1); +x_7 = lean_ptr_addr(x_4); +x_8 = lean_usize_to_uint64(x_7); +x_9 = lean_hashmap_mk_idx(x_6, x_8); +x_10 = lean_array_uget(x_1, x_9); +lean_ctor_set(x_2, 2, x_10); +x_11 = lean_array_uset(x_1, x_9, x_2); +x_1 = x_11; +x_2 = x_5; +goto _start; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; size_t x_17; uint64_t x_18; size_t x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_13 = lean_ctor_get(x_2, 0); +x_14 = lean_ctor_get(x_2, 1); +x_15 = lean_ctor_get(x_2, 2); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_dec(x_2); +x_16 = lean_array_get_size(x_1); +x_17 = lean_ptr_addr(x_13); +x_18 = lean_usize_to_uint64(x_17); +x_19 = lean_hashmap_mk_idx(x_16, x_18); +x_20 = lean_array_uget(x_1, x_19); +x_21 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_21, 0, x_13); +lean_ctor_set(x_21, 1, x_14); +lean_ctor_set(x_21, 2, x_20); +x_22 = lean_array_uset(x_1, x_19, x_21); +x_1 = x_22; +x_2 = x_15; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_HashMapImp_moveEntries___at_Lean_Meta_Canonicalizer_canon_unsafe__2___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_array_get_size(x_2); +x_5 = lean_nat_dec_lt(x_1, x_4); +lean_dec(x_4); +if (x_5 == 0) +{ +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_6 = lean_array_fget(x_2, x_1); +x_7 = lean_box(0); +x_8 = lean_array_fset(x_2, x_1, x_7); +x_9 = l_Lean_AssocList_foldlM___at_Lean_Meta_Canonicalizer_canon_unsafe__2___spec__4___at_Lean_Meta_Canonicalizer_canon_unsafe__2___spec__5(x_3, x_6); +x_10 = lean_unsigned_to_nat(1u); +x_11 = lean_nat_add(x_1, x_10); +lean_dec(x_1); +x_1 = x_11; +x_2 = x_8; +x_3 = x_9; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Lean_HashMapImp_expand___at_Lean_Meta_Canonicalizer_canon_unsafe__2___spec__2(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_3 = lean_array_get_size(x_2); +x_4 = lean_unsigned_to_nat(2u); +x_5 = lean_nat_mul(x_3, x_4); +lean_dec(x_3); +x_6 = lean_box(0); +x_7 = lean_mk_array(x_5, x_6); +x_8 = lean_unsigned_to_nat(0u); +x_9 = l_Lean_HashMapImp_moveEntries___at_Lean_Meta_Canonicalizer_canon_unsafe__2___spec__3(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_Meta_Canonicalizer_canon_unsafe__2___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_4; +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(0); +return x_4; +} +else +{ +uint8_t x_5; +x_5 = !lean_is_exclusive(x_3); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; uint8_t x_11; +x_6 = lean_ctor_get(x_3, 0); +x_7 = lean_ctor_get(x_3, 1); +x_8 = lean_ctor_get(x_3, 2); +x_9 = lean_ptr_addr(x_6); +x_10 = lean_ptr_addr(x_1); +x_11 = lean_usize_dec_eq(x_9, x_10); +if (x_11 == 0) +{ +lean_object* x_12; +x_12 = l_Lean_AssocList_replace___at_Lean_Meta_Canonicalizer_canon_unsafe__2___spec__6(x_1, x_2, x_8); +lean_ctor_set(x_3, 2, x_12); +return x_3; +} +else +{ +lean_dec(x_7); +lean_dec(x_6); +lean_ctor_set(x_3, 1, x_2); +lean_ctor_set(x_3, 0, x_1); +return x_3; +} +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; size_t x_16; size_t x_17; uint8_t x_18; +x_13 = lean_ctor_get(x_3, 0); +x_14 = lean_ctor_get(x_3, 1); +x_15 = lean_ctor_get(x_3, 2); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_dec(x_3); +x_16 = lean_ptr_addr(x_13); +x_17 = lean_ptr_addr(x_1); +x_18 = lean_usize_dec_eq(x_16, x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +x_19 = l_Lean_AssocList_replace___at_Lean_Meta_Canonicalizer_canon_unsafe__2___spec__6(x_1, x_2, x_15); +x_20 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_20, 0, x_13); +lean_ctor_set(x_20, 1, x_14); +lean_ctor_set(x_20, 2, x_19); +return x_20; +} +else +{ +lean_object* x_21; +lean_dec(x_14); +lean_dec(x_13); +x_21 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_21, 0, x_1); +lean_ctor_set(x_21, 1, x_2); +lean_ctor_set(x_21, 2, x_15); +return x_21; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Canonicalizer_canon_unsafe__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; uint8_t x_14; +x_11 = lean_st_ref_take(x_5, 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; uint8_t x_17; +x_15 = lean_ctor_get(x_12, 2); +x_16 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_16, 0, x_1); +lean_ctor_set(x_16, 1, x_3); +x_17 = !lean_is_exclusive(x_15); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; size_t x_21; uint64_t x_22; size_t x_23; lean_object* x_24; uint8_t x_25; +x_18 = lean_ctor_get(x_15, 0); +x_19 = lean_ctor_get(x_15, 1); +x_20 = lean_array_get_size(x_19); +x_21 = lean_ptr_addr(x_2); +x_22 = lean_usize_to_uint64(x_21); +lean_inc(x_20); +x_23 = lean_hashmap_mk_idx(x_20, x_22); +x_24 = lean_array_uget(x_19, x_23); +x_25 = l_Lean_AssocList_contains___at_Lean_Meta_Canonicalizer_canon_unsafe__2___spec__1(x_2, x_24); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; +x_26 = lean_unsigned_to_nat(1u); +x_27 = lean_nat_add(x_18, x_26); +lean_dec(x_18); +x_28 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_28, 0, x_2); +lean_ctor_set(x_28, 1, x_16); +lean_ctor_set(x_28, 2, x_24); +x_29 = lean_array_uset(x_19, x_23, x_28); +x_30 = l___private_Lean_Data_HashMap_0__Lean_numBucketsForCapacity(x_27); +x_31 = lean_nat_dec_le(x_30, x_20); +lean_dec(x_20); +lean_dec(x_30); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; uint8_t x_34; +lean_free_object(x_15); +x_32 = l_Lean_HashMapImp_expand___at_Lean_Meta_Canonicalizer_canon_unsafe__2___spec__2(x_27, x_29); +lean_ctor_set(x_12, 2, x_32); +x_33 = lean_st_ref_set(x_5, x_12, x_13); +x_34 = !lean_is_exclusive(x_33); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; +x_35 = lean_ctor_get(x_33, 0); +lean_dec(x_35); +x_36 = lean_box(0); +lean_ctor_set(x_33, 0, x_36); +return x_33; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_33, 1); +lean_inc(x_37); +lean_dec(x_33); +x_38 = lean_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_37); +return x_39; +} +} +else +{ +lean_object* x_40; uint8_t x_41; +lean_ctor_set(x_15, 1, x_29); +lean_ctor_set(x_15, 0, x_27); +x_40 = lean_st_ref_set(x_5, x_12, x_13); +x_41 = !lean_is_exclusive(x_40); +if (x_41 == 0) +{ +lean_object* x_42; lean_object* x_43; +x_42 = lean_ctor_get(x_40, 0); +lean_dec(x_42); +x_43 = lean_box(0); +lean_ctor_set(x_40, 0, x_43); +return x_40; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_40, 1); +lean_inc(x_44); +lean_dec(x_40); +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 +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; uint8_t x_50; +lean_dec(x_20); +x_47 = l_Lean_AssocList_replace___at_Lean_Meta_Canonicalizer_canon_unsafe__2___spec__6(x_2, x_16, x_24); +x_48 = lean_array_uset(x_19, x_23, x_47); +lean_ctor_set(x_15, 1, x_48); +x_49 = lean_st_ref_set(x_5, x_12, x_13); +x_50 = !lean_is_exclusive(x_49); +if (x_50 == 0) +{ +lean_object* x_51; lean_object* x_52; +x_51 = lean_ctor_get(x_49, 0); +lean_dec(x_51); +x_52 = lean_box(0); +lean_ctor_set(x_49, 0, x_52); +return x_49; +} +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 = 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; +} +} +} +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; size_t x_59; uint64_t x_60; size_t x_61; lean_object* x_62; uint8_t x_63; +x_56 = lean_ctor_get(x_15, 0); +x_57 = lean_ctor_get(x_15, 1); +lean_inc(x_57); +lean_inc(x_56); +lean_dec(x_15); +x_58 = lean_array_get_size(x_57); +x_59 = lean_ptr_addr(x_2); +x_60 = lean_usize_to_uint64(x_59); +lean_inc(x_58); +x_61 = lean_hashmap_mk_idx(x_58, x_60); +x_62 = lean_array_uget(x_57, x_61); +x_63 = l_Lean_AssocList_contains___at_Lean_Meta_Canonicalizer_canon_unsafe__2___spec__1(x_2, x_62); +if (x_63 == 0) +{ +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_64 = lean_unsigned_to_nat(1u); +x_65 = lean_nat_add(x_56, x_64); +lean_dec(x_56); +x_66 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_66, 0, x_2); +lean_ctor_set(x_66, 1, x_16); +lean_ctor_set(x_66, 2, x_62); +x_67 = lean_array_uset(x_57, x_61, x_66); +x_68 = l___private_Lean_Data_HashMap_0__Lean_numBucketsForCapacity(x_65); +x_69 = lean_nat_dec_le(x_68, x_58); +lean_dec(x_58); +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; +x_70 = l_Lean_HashMapImp_expand___at_Lean_Meta_Canonicalizer_canon_unsafe__2___spec__2(x_65, x_67); +lean_ctor_set(x_12, 2, x_70); +x_71 = lean_st_ref_set(x_5, x_12, x_13); +x_72 = lean_ctor_get(x_71, 1); +lean_inc(x_72); +if (lean_is_exclusive(x_71)) { + lean_ctor_release(x_71, 0); + lean_ctor_release(x_71, 1); + x_73 = x_71; +} else { + lean_dec_ref(x_71); + x_73 = lean_box(0); +} +x_74 = lean_box(0); +if (lean_is_scalar(x_73)) { + x_75 = lean_alloc_ctor(0, 2, 0); +} else { + x_75 = x_73; +} +lean_ctor_set(x_75, 0, x_74); +lean_ctor_set(x_75, 1, x_72); +return x_75; +} +else +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_76 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_76, 0, x_65); +lean_ctor_set(x_76, 1, x_67); +lean_ctor_set(x_12, 2, x_76); +x_77 = lean_st_ref_set(x_5, x_12, x_13); +x_78 = lean_ctor_get(x_77, 1); +lean_inc(x_78); +if (lean_is_exclusive(x_77)) { + lean_ctor_release(x_77, 0); + lean_ctor_release(x_77, 1); + x_79 = x_77; +} else { + lean_dec_ref(x_77); + x_79 = lean_box(0); +} +x_80 = lean_box(0); +if (lean_is_scalar(x_79)) { + x_81 = lean_alloc_ctor(0, 2, 0); +} else { + x_81 = x_79; +} +lean_ctor_set(x_81, 0, x_80); +lean_ctor_set(x_81, 1, x_78); +return x_81; +} +} +else +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; +lean_dec(x_58); +x_82 = l_Lean_AssocList_replace___at_Lean_Meta_Canonicalizer_canon_unsafe__2___spec__6(x_2, x_16, x_62); +x_83 = lean_array_uset(x_57, x_61, x_82); +x_84 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_84, 0, x_56); +lean_ctor_set(x_84, 1, x_83); +lean_ctor_set(x_12, 2, x_84); +x_85 = lean_st_ref_set(x_5, x_12, x_13); +x_86 = lean_ctor_get(x_85, 1); +lean_inc(x_86); +if (lean_is_exclusive(x_85)) { + lean_ctor_release(x_85, 0); + lean_ctor_release(x_85, 1); + x_87 = x_85; +} else { + lean_dec_ref(x_85); + x_87 = lean_box(0); +} +x_88 = lean_box(0); +if (lean_is_scalar(x_87)) { + x_89 = lean_alloc_ctor(0, 2, 0); +} else { + x_89 = x_87; +} +lean_ctor_set(x_89, 0, x_88); +lean_ctor_set(x_89, 1, x_86); +return x_89; +} +} +} +else +{ +lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; size_t x_98; uint64_t x_99; size_t x_100; lean_object* x_101; uint8_t x_102; +x_90 = lean_ctor_get(x_12, 0); +x_91 = lean_ctor_get(x_12, 1); +x_92 = lean_ctor_get(x_12, 2); +lean_inc(x_92); +lean_inc(x_91); +lean_inc(x_90); +lean_dec(x_12); +x_93 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_93, 0, x_1); +lean_ctor_set(x_93, 1, x_3); +x_94 = lean_ctor_get(x_92, 0); +lean_inc(x_94); +x_95 = lean_ctor_get(x_92, 1); +lean_inc(x_95); +if (lean_is_exclusive(x_92)) { + lean_ctor_release(x_92, 0); + lean_ctor_release(x_92, 1); + x_96 = x_92; +} else { + lean_dec_ref(x_92); + x_96 = lean_box(0); +} +x_97 = lean_array_get_size(x_95); +x_98 = lean_ptr_addr(x_2); +x_99 = lean_usize_to_uint64(x_98); +lean_inc(x_97); +x_100 = lean_hashmap_mk_idx(x_97, x_99); +x_101 = lean_array_uget(x_95, x_100); +x_102 = l_Lean_AssocList_contains___at_Lean_Meta_Canonicalizer_canon_unsafe__2___spec__1(x_2, x_101); +if (x_102 == 0) +{ +lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; uint8_t x_108; +x_103 = lean_unsigned_to_nat(1u); +x_104 = lean_nat_add(x_94, x_103); +lean_dec(x_94); +x_105 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_105, 0, x_2); +lean_ctor_set(x_105, 1, x_93); +lean_ctor_set(x_105, 2, x_101); +x_106 = lean_array_uset(x_95, x_100, x_105); +x_107 = l___private_Lean_Data_HashMap_0__Lean_numBucketsForCapacity(x_104); +x_108 = lean_nat_dec_le(x_107, x_97); +lean_dec(x_97); +lean_dec(x_107); +if (x_108 == 0) +{ +lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; +lean_dec(x_96); +x_109 = l_Lean_HashMapImp_expand___at_Lean_Meta_Canonicalizer_canon_unsafe__2___spec__2(x_104, x_106); +x_110 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_110, 0, x_90); +lean_ctor_set(x_110, 1, x_91); +lean_ctor_set(x_110, 2, x_109); +x_111 = lean_st_ref_set(x_5, x_110, x_13); +x_112 = lean_ctor_get(x_111, 1); +lean_inc(x_112); +if (lean_is_exclusive(x_111)) { + lean_ctor_release(x_111, 0); + lean_ctor_release(x_111, 1); + x_113 = x_111; +} else { + lean_dec_ref(x_111); + x_113 = lean_box(0); +} +x_114 = lean_box(0); +if (lean_is_scalar(x_113)) { + x_115 = lean_alloc_ctor(0, 2, 0); +} else { + x_115 = x_113; +} +lean_ctor_set(x_115, 0, x_114); +lean_ctor_set(x_115, 1, x_112); +return x_115; +} +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; +if (lean_is_scalar(x_96)) { + x_116 = lean_alloc_ctor(0, 2, 0); +} else { + x_116 = x_96; +} +lean_ctor_set(x_116, 0, x_104); +lean_ctor_set(x_116, 1, x_106); +x_117 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_117, 0, x_90); +lean_ctor_set(x_117, 1, x_91); +lean_ctor_set(x_117, 2, x_116); +x_118 = lean_st_ref_set(x_5, x_117, x_13); +x_119 = lean_ctor_get(x_118, 1); +lean_inc(x_119); +if (lean_is_exclusive(x_118)) { + lean_ctor_release(x_118, 0); + lean_ctor_release(x_118, 1); + x_120 = x_118; +} else { + lean_dec_ref(x_118); + x_120 = lean_box(0); +} +x_121 = lean_box(0); +if (lean_is_scalar(x_120)) { + x_122 = lean_alloc_ctor(0, 2, 0); +} else { + x_122 = x_120; +} +lean_ctor_set(x_122, 0, x_121); +lean_ctor_set(x_122, 1, x_119); +return x_122; +} +} +else +{ +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_dec(x_97); +x_123 = l_Lean_AssocList_replace___at_Lean_Meta_Canonicalizer_canon_unsafe__2___spec__6(x_2, x_93, x_101); +x_124 = lean_array_uset(x_95, x_100, x_123); +if (lean_is_scalar(x_96)) { + x_125 = lean_alloc_ctor(0, 2, 0); +} else { + x_125 = x_96; +} +lean_ctor_set(x_125, 0, x_94); +lean_ctor_set(x_125, 1, x_124); +x_126 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_126, 0, x_90); +lean_ctor_set(x_126, 1, x_91); +lean_ctor_set(x_126, 2, x_125); +x_127 = lean_st_ref_set(x_5, x_126, x_13); +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_box(0); +if (lean_is_scalar(x_129)) { + x_131 = lean_alloc_ctor(0, 2, 0); +} else { + x_131 = x_129; +} +lean_ctor_set(x_131, 0, x_130); +lean_ctor_set(x_131, 1, x_128); +return x_131; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_AssocList_contains___at_Lean_Meta_Canonicalizer_canon_unsafe__2___spec__1___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_Meta_Canonicalizer_canon_unsafe__2___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_Lean_Meta_Canonicalizer_canon_unsafe__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_Meta_Canonicalizer_canon_unsafe__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_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Canonicalizer_canon_unsafe__3(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_10 = lean_st_ref_take(x_4, x_9); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +x_13 = !lean_is_exclusive(x_11); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_14 = lean_ctor_get(x_11, 2); +x_15 = lean_box(0); +x_16 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_16, 0, x_1); +lean_ctor_set(x_16, 1, x_15); +x_17 = !lean_is_exclusive(x_14); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; size_t x_21; uint64_t x_22; size_t x_23; lean_object* x_24; uint8_t x_25; +x_18 = lean_ctor_get(x_14, 0); +x_19 = lean_ctor_get(x_14, 1); +x_20 = lean_array_get_size(x_19); +x_21 = lean_ptr_addr(x_2); +x_22 = lean_usize_to_uint64(x_21); +lean_inc(x_20); +x_23 = lean_hashmap_mk_idx(x_20, x_22); +x_24 = lean_array_uget(x_19, x_23); +x_25 = l_Lean_AssocList_contains___at_Lean_Meta_Canonicalizer_canon_unsafe__2___spec__1(x_2, x_24); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; +x_26 = lean_unsigned_to_nat(1u); +x_27 = lean_nat_add(x_18, x_26); +lean_dec(x_18); +x_28 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_28, 0, x_2); +lean_ctor_set(x_28, 1, x_16); +lean_ctor_set(x_28, 2, x_24); +x_29 = lean_array_uset(x_19, x_23, x_28); +x_30 = l___private_Lean_Data_HashMap_0__Lean_numBucketsForCapacity(x_27); +x_31 = lean_nat_dec_le(x_30, x_20); +lean_dec(x_20); +lean_dec(x_30); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; uint8_t x_34; +lean_free_object(x_14); +x_32 = l_Lean_HashMapImp_expand___at_Lean_Meta_Canonicalizer_canon_unsafe__2___spec__2(x_27, x_29); +lean_ctor_set(x_11, 2, x_32); +x_33 = lean_st_ref_set(x_4, x_11, x_12); +x_34 = !lean_is_exclusive(x_33); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; +x_35 = lean_ctor_get(x_33, 0); +lean_dec(x_35); +x_36 = lean_box(0); +lean_ctor_set(x_33, 0, x_36); +return x_33; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_33, 1); +lean_inc(x_37); +lean_dec(x_33); +x_38 = lean_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_37); +return x_39; +} +} +else +{ +lean_object* x_40; uint8_t x_41; +lean_ctor_set(x_14, 1, x_29); +lean_ctor_set(x_14, 0, x_27); +x_40 = lean_st_ref_set(x_4, x_11, x_12); +x_41 = !lean_is_exclusive(x_40); +if (x_41 == 0) +{ +lean_object* x_42; lean_object* x_43; +x_42 = lean_ctor_get(x_40, 0); +lean_dec(x_42); +x_43 = lean_box(0); +lean_ctor_set(x_40, 0, x_43); +return x_40; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_40, 1); +lean_inc(x_44); +lean_dec(x_40); +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 +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; uint8_t x_50; +lean_dec(x_20); +x_47 = l_Lean_AssocList_replace___at_Lean_Meta_Canonicalizer_canon_unsafe__2___spec__6(x_2, x_16, x_24); +x_48 = lean_array_uset(x_19, x_23, x_47); +lean_ctor_set(x_14, 1, x_48); +x_49 = lean_st_ref_set(x_4, x_11, x_12); +x_50 = !lean_is_exclusive(x_49); +if (x_50 == 0) +{ +lean_object* x_51; lean_object* x_52; +x_51 = lean_ctor_get(x_49, 0); +lean_dec(x_51); +x_52 = lean_box(0); +lean_ctor_set(x_49, 0, x_52); +return x_49; +} +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 = 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; +} +} +} +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; size_t x_59; uint64_t x_60; size_t x_61; lean_object* x_62; uint8_t x_63; +x_56 = lean_ctor_get(x_14, 0); +x_57 = lean_ctor_get(x_14, 1); +lean_inc(x_57); +lean_inc(x_56); +lean_dec(x_14); +x_58 = lean_array_get_size(x_57); +x_59 = lean_ptr_addr(x_2); +x_60 = lean_usize_to_uint64(x_59); +lean_inc(x_58); +x_61 = lean_hashmap_mk_idx(x_58, x_60); +x_62 = lean_array_uget(x_57, x_61); +x_63 = l_Lean_AssocList_contains___at_Lean_Meta_Canonicalizer_canon_unsafe__2___spec__1(x_2, x_62); +if (x_63 == 0) +{ +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_64 = lean_unsigned_to_nat(1u); +x_65 = lean_nat_add(x_56, x_64); +lean_dec(x_56); +x_66 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_66, 0, x_2); +lean_ctor_set(x_66, 1, x_16); +lean_ctor_set(x_66, 2, x_62); +x_67 = lean_array_uset(x_57, x_61, x_66); +x_68 = l___private_Lean_Data_HashMap_0__Lean_numBucketsForCapacity(x_65); +x_69 = lean_nat_dec_le(x_68, x_58); +lean_dec(x_58); +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; +x_70 = l_Lean_HashMapImp_expand___at_Lean_Meta_Canonicalizer_canon_unsafe__2___spec__2(x_65, x_67); +lean_ctor_set(x_11, 2, x_70); +x_71 = lean_st_ref_set(x_4, x_11, x_12); +x_72 = lean_ctor_get(x_71, 1); +lean_inc(x_72); +if (lean_is_exclusive(x_71)) { + lean_ctor_release(x_71, 0); + lean_ctor_release(x_71, 1); + x_73 = x_71; +} else { + lean_dec_ref(x_71); + x_73 = lean_box(0); +} +x_74 = lean_box(0); +if (lean_is_scalar(x_73)) { + x_75 = lean_alloc_ctor(0, 2, 0); +} else { + x_75 = x_73; +} +lean_ctor_set(x_75, 0, x_74); +lean_ctor_set(x_75, 1, x_72); +return x_75; +} +else +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_76 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_76, 0, x_65); +lean_ctor_set(x_76, 1, x_67); +lean_ctor_set(x_11, 2, x_76); +x_77 = lean_st_ref_set(x_4, x_11, x_12); +x_78 = lean_ctor_get(x_77, 1); +lean_inc(x_78); +if (lean_is_exclusive(x_77)) { + lean_ctor_release(x_77, 0); + lean_ctor_release(x_77, 1); + x_79 = x_77; +} else { + lean_dec_ref(x_77); + x_79 = lean_box(0); +} +x_80 = lean_box(0); +if (lean_is_scalar(x_79)) { + x_81 = lean_alloc_ctor(0, 2, 0); +} else { + x_81 = x_79; +} +lean_ctor_set(x_81, 0, x_80); +lean_ctor_set(x_81, 1, x_78); +return x_81; +} +} +else +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; +lean_dec(x_58); +x_82 = l_Lean_AssocList_replace___at_Lean_Meta_Canonicalizer_canon_unsafe__2___spec__6(x_2, x_16, x_62); +x_83 = lean_array_uset(x_57, x_61, x_82); +x_84 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_84, 0, x_56); +lean_ctor_set(x_84, 1, x_83); +lean_ctor_set(x_11, 2, x_84); +x_85 = lean_st_ref_set(x_4, x_11, x_12); +x_86 = lean_ctor_get(x_85, 1); +lean_inc(x_86); +if (lean_is_exclusive(x_85)) { + lean_ctor_release(x_85, 0); + lean_ctor_release(x_85, 1); + x_87 = x_85; +} else { + lean_dec_ref(x_85); + x_87 = lean_box(0); +} +x_88 = lean_box(0); +if (lean_is_scalar(x_87)) { + x_89 = lean_alloc_ctor(0, 2, 0); +} else { + x_89 = x_87; +} +lean_ctor_set(x_89, 0, x_88); +lean_ctor_set(x_89, 1, x_86); +return x_89; +} +} +} +else +{ +lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; size_t x_99; uint64_t x_100; size_t x_101; lean_object* x_102; uint8_t x_103; +x_90 = lean_ctor_get(x_11, 0); +x_91 = lean_ctor_get(x_11, 1); +x_92 = lean_ctor_get(x_11, 2); +lean_inc(x_92); +lean_inc(x_91); +lean_inc(x_90); +lean_dec(x_11); +x_93 = lean_box(0); +x_94 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_94, 0, x_1); +lean_ctor_set(x_94, 1, x_93); +x_95 = lean_ctor_get(x_92, 0); +lean_inc(x_95); +x_96 = lean_ctor_get(x_92, 1); +lean_inc(x_96); +if (lean_is_exclusive(x_92)) { + lean_ctor_release(x_92, 0); + lean_ctor_release(x_92, 1); + x_97 = x_92; +} else { + lean_dec_ref(x_92); + x_97 = lean_box(0); +} +x_98 = lean_array_get_size(x_96); +x_99 = lean_ptr_addr(x_2); +x_100 = lean_usize_to_uint64(x_99); +lean_inc(x_98); +x_101 = lean_hashmap_mk_idx(x_98, x_100); +x_102 = lean_array_uget(x_96, x_101); +x_103 = l_Lean_AssocList_contains___at_Lean_Meta_Canonicalizer_canon_unsafe__2___spec__1(x_2, x_102); +if (x_103 == 0) +{ +lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; uint8_t x_109; +x_104 = lean_unsigned_to_nat(1u); +x_105 = lean_nat_add(x_95, x_104); +lean_dec(x_95); +x_106 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_106, 0, x_2); +lean_ctor_set(x_106, 1, x_94); +lean_ctor_set(x_106, 2, x_102); +x_107 = lean_array_uset(x_96, x_101, x_106); +x_108 = l___private_Lean_Data_HashMap_0__Lean_numBucketsForCapacity(x_105); +x_109 = lean_nat_dec_le(x_108, x_98); +lean_dec(x_98); +lean_dec(x_108); +if (x_109 == 0) +{ +lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +lean_dec(x_97); +x_110 = l_Lean_HashMapImp_expand___at_Lean_Meta_Canonicalizer_canon_unsafe__2___spec__2(x_105, x_107); +x_111 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_111, 0, x_90); +lean_ctor_set(x_111, 1, x_91); +lean_ctor_set(x_111, 2, x_110); +x_112 = lean_st_ref_set(x_4, x_111, x_12); +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); +} +x_115 = lean_box(0); +if (lean_is_scalar(x_114)) { + x_116 = lean_alloc_ctor(0, 2, 0); +} else { + x_116 = x_114; +} +lean_ctor_set(x_116, 0, x_115); +lean_ctor_set(x_116, 1, x_113); +return x_116; +} +else +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; +if (lean_is_scalar(x_97)) { + x_117 = lean_alloc_ctor(0, 2, 0); +} else { + x_117 = x_97; +} +lean_ctor_set(x_117, 0, x_105); +lean_ctor_set(x_117, 1, x_107); +x_118 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_118, 0, x_90); +lean_ctor_set(x_118, 1, x_91); +lean_ctor_set(x_118, 2, x_117); +x_119 = lean_st_ref_set(x_4, x_118, x_12); +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 +{ +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_dec(x_98); +x_124 = l_Lean_AssocList_replace___at_Lean_Meta_Canonicalizer_canon_unsafe__2___spec__6(x_2, x_94, x_102); +x_125 = lean_array_uset(x_96, x_101, x_124); +if (lean_is_scalar(x_97)) { + x_126 = lean_alloc_ctor(0, 2, 0); +} else { + x_126 = x_97; +} +lean_ctor_set(x_126, 0, x_95); +lean_ctor_set(x_126, 1, x_125); +x_127 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_127, 0, x_90); +lean_ctor_set(x_127, 1, x_91); +lean_ctor_set(x_127, 2, x_126); +x_128 = lean_st_ref_set(x_4, x_127, x_12); +x_129 = lean_ctor_get(x_128, 1); +lean_inc(x_129); +if (lean_is_exclusive(x_128)) { + lean_ctor_release(x_128, 0); + lean_ctor_release(x_128, 1); + x_130 = x_128; +} else { + lean_dec_ref(x_128); + x_130 = lean_box(0); +} +x_131 = lean_box(0); +if (lean_is_scalar(x_130)) { + x_132 = lean_alloc_ctor(0, 2, 0); +} else { + x_132 = x_130; +} +lean_ctor_set(x_132, 0, x_131); +lean_ctor_set(x_132, 1, x_129); +return x_132; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Canonicalizer_canon_unsafe__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: +{ +uint8_t x_10; lean_object* x_11; +x_10 = lean_unbox(x_3); +lean_dec(x_3); +x_11 = l_Lean_Meta_Canonicalizer_canon_unsafe__3(x_1, x_2, x_10, 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_11; +} +} +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Meta_Canonicalizer_canon___spec__1(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) { +_start: +{ +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_12; +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_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_4); +lean_ctor_set(x_12, 1, x_11); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +lean_dec(x_4); +x_13 = lean_ctor_get(x_3, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_3, 1); +lean_inc(x_14); +lean_dec(x_3); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_13); +lean_inc(x_1); +x_15 = l_Lean_Meta_isExprDefEq(x_1, x_13, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; uint8_t x_17; +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_18; +lean_dec(x_13); +x_18 = lean_ctor_get(x_15, 1); +lean_inc(x_18); +lean_dec(x_15); +lean_inc(x_2); +{ +lean_object* _tmp_2 = x_14; +lean_object* _tmp_3 = x_2; +lean_object* _tmp_10 = x_18; +x_3 = _tmp_2; +x_4 = _tmp_3; +x_11 = _tmp_10; +} +goto _start; +} +else +{ +uint8_t x_20; +lean_dec(x_14); +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_20 = !lean_is_exclusive(x_15); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_21 = lean_ctor_get(x_15, 0); +lean_dec(x_21); +x_22 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_22, 0, x_13); +x_23 = lean_box(0); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +lean_ctor_set(x_15, 0, x_24); +return x_15; +} +else +{ +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_15, 1); +lean_inc(x_25); +lean_dec(x_15); +x_26 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_26, 0, x_13); +x_27 = lean_box(0); +x_28 = lean_alloc_ctor(0, 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_25); +return x_29; +} +} +} +else +{ +uint8_t x_30; +lean_dec(x_14); +lean_dec(x_13); +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_30 = !lean_is_exclusive(x_15); +if (x_30 == 0) +{ +return x_15; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_15, 0); +x_32 = lean_ctor_get(x_15, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_15); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Canonicalizer_canon___lambda__1(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) { +_start: +{ +lean_object* x_12; uint8_t x_13; +lean_inc(x_1); +x_12 = l_Lean_Meta_Canonicalizer_canon_unsafe__2(x_1, x_2, x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +lean_object* x_14; +x_14 = lean_ctor_get(x_12, 0); +lean_dec(x_14); +lean_ctor_set(x_12, 0, x_1); +return x_12; +} +else +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_12, 1); +lean_inc(x_15); +lean_dec(x_12); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_1); +lean_ctor_set(x_16, 1, x_15); +return x_16; +} +} +} +static lean_object* _init_l_Lean_Meta_Canonicalizer_canon___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +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); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Canonicalizer_canon(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_9; +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_9 = l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey(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; 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_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_3, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +lean_inc(x_10); +x_15 = l_Lean_Meta_Canonicalizer_canon_unsafe__1(x_10, x_13); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; uint8_t x_17; +lean_inc(x_1); +x_16 = l_Lean_Meta_Canonicalizer_canon_unsafe__3(x_1, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_14); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +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_1); +return x_16; +} +else +{ +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_1); +lean_ctor_set(x_20, 1, x_19); +return x_20; +} +} +else +{ +lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_21 = lean_ctor_get(x_4, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_15, 0); +lean_inc(x_22); +lean_dec(x_15); +x_23 = !lean_is_exclusive(x_4); +if (x_23 == 0) +{ +lean_object* x_24; uint8_t x_25; +x_24 = lean_ctor_get(x_4, 0); +lean_dec(x_24); +x_25 = !lean_is_exclusive(x_21); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; +lean_ctor_set_uint8(x_21, 9, x_2); +x_26 = l_Lean_Meta_Canonicalizer_canon___closed__1; +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_22); +lean_inc(x_1); +x_27 = l_List_forIn_loop___at_Lean_Meta_Canonicalizer_canon___spec__1(x_1, x_26, x_22, x_26, x_2, x_3, x_4, x_5, x_6, x_7, x_14); +if (lean_obj_tag(x_27) == 0) +{ +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +lean_dec(x_28); +if (lean_obj_tag(x_29) == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +x_30 = lean_ctor_get(x_27, 1); +lean_inc(x_30); +lean_dec(x_27); +x_31 = lean_box(0); +x_32 = l_Lean_Meta_Canonicalizer_canon___lambda__1(x_1, x_10, x_22, x_31, x_2, x_3, x_4, x_5, x_6, x_7, x_30); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_33 = !lean_is_exclusive(x_32); +if (x_33 == 0) +{ +return x_32; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_32, 0); +x_35 = lean_ctor_get(x_32, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_32); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; +} +} +else +{ +uint8_t x_37; +lean_dec(x_4); +lean_dec(x_22); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_37 = !lean_is_exclusive(x_27); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_27, 0); +lean_dec(x_38); +x_39 = lean_ctor_get(x_29, 0); +lean_inc(x_39); +lean_dec(x_29); +lean_ctor_set(x_27, 0, x_39); +return x_27; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_27, 1); +lean_inc(x_40); +lean_dec(x_27); +x_41 = lean_ctor_get(x_29, 0); +lean_inc(x_41); +lean_dec(x_29); +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; +} +} +} +else +{ +uint8_t x_43; +lean_dec(x_4); +lean_dec(x_22); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_27); +if (x_43 == 0) +{ +return x_27; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_27, 0); +x_45 = lean_ctor_get(x_27, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_27); +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; +} +} +} +else +{ +uint8_t x_47; uint8_t x_48; uint8_t x_49; uint8_t 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; lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_47 = lean_ctor_get_uint8(x_21, 0); +x_48 = lean_ctor_get_uint8(x_21, 1); +x_49 = lean_ctor_get_uint8(x_21, 2); +x_50 = lean_ctor_get_uint8(x_21, 3); +x_51 = lean_ctor_get_uint8(x_21, 4); +x_52 = lean_ctor_get_uint8(x_21, 5); +x_53 = lean_ctor_get_uint8(x_21, 6); +x_54 = lean_ctor_get_uint8(x_21, 7); +x_55 = lean_ctor_get_uint8(x_21, 8); +x_56 = lean_ctor_get_uint8(x_21, 10); +x_57 = lean_ctor_get_uint8(x_21, 11); +lean_dec(x_21); +x_58 = lean_alloc_ctor(0, 0, 12); +lean_ctor_set_uint8(x_58, 0, x_47); +lean_ctor_set_uint8(x_58, 1, x_48); +lean_ctor_set_uint8(x_58, 2, x_49); +lean_ctor_set_uint8(x_58, 3, x_50); +lean_ctor_set_uint8(x_58, 4, x_51); +lean_ctor_set_uint8(x_58, 5, x_52); +lean_ctor_set_uint8(x_58, 6, x_53); +lean_ctor_set_uint8(x_58, 7, x_54); +lean_ctor_set_uint8(x_58, 8, x_55); +lean_ctor_set_uint8(x_58, 9, x_2); +lean_ctor_set_uint8(x_58, 10, x_56); +lean_ctor_set_uint8(x_58, 11, x_57); +lean_ctor_set(x_4, 0, x_58); +x_59 = l_Lean_Meta_Canonicalizer_canon___closed__1; +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_22); +lean_inc(x_1); +x_60 = l_List_forIn_loop___at_Lean_Meta_Canonicalizer_canon___spec__1(x_1, x_59, x_22, x_59, x_2, x_3, x_4, x_5, x_6, x_7, x_14); +if (lean_obj_tag(x_60) == 0) +{ +lean_object* x_61; lean_object* x_62; +x_61 = lean_ctor_get(x_60, 0); +lean_inc(x_61); +x_62 = lean_ctor_get(x_61, 0); +lean_inc(x_62); +lean_dec(x_61); +if (lean_obj_tag(x_62) == 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_60, 1); +lean_inc(x_63); +lean_dec(x_60); +x_64 = lean_box(0); +x_65 = l_Lean_Meta_Canonicalizer_canon___lambda__1(x_1, x_10, x_22, x_64, x_2, x_3, x_4, x_5, x_6, x_7, x_63); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_66 = lean_ctor_get(x_65, 0); +lean_inc(x_66); +x_67 = lean_ctor_get(x_65, 1); +lean_inc(x_67); +if (lean_is_exclusive(x_65)) { + lean_ctor_release(x_65, 0); + lean_ctor_release(x_65, 1); + x_68 = x_65; +} else { + lean_dec_ref(x_65); + 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_66); +lean_ctor_set(x_69, 1, x_67); +return x_69; +} +else +{ +lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; +lean_dec(x_4); +lean_dec(x_22); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_70 = lean_ctor_get(x_60, 1); +lean_inc(x_70); +if (lean_is_exclusive(x_60)) { + lean_ctor_release(x_60, 0); + lean_ctor_release(x_60, 1); + x_71 = x_60; +} else { + lean_dec_ref(x_60); + x_71 = lean_box(0); +} +x_72 = lean_ctor_get(x_62, 0); +lean_inc(x_72); +lean_dec(x_62); +if (lean_is_scalar(x_71)) { + x_73 = lean_alloc_ctor(0, 2, 0); +} else { + x_73 = 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_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; +lean_dec(x_4); +lean_dec(x_22); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_74 = lean_ctor_get(x_60, 0); +lean_inc(x_74); +x_75 = lean_ctor_get(x_60, 1); +lean_inc(x_75); +if (lean_is_exclusive(x_60)) { + lean_ctor_release(x_60, 0); + lean_ctor_release(x_60, 1); + x_76 = x_60; +} else { + lean_dec_ref(x_60); + x_76 = lean_box(0); +} +if (lean_is_scalar(x_76)) { + x_77 = lean_alloc_ctor(1, 2, 0); +} else { + x_77 = x_76; +} +lean_ctor_set(x_77, 0, x_74); +lean_ctor_set(x_77, 1, x_75); +return x_77; +} +} +} +else +{ +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; uint8_t x_83; uint8_t x_84; uint8_t x_85; uint8_t x_86; uint8_t x_87; uint8_t x_88; uint8_t x_89; uint8_t x_90; uint8_t x_91; uint8_t x_92; uint8_t x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_78 = lean_ctor_get(x_4, 1); +x_79 = lean_ctor_get(x_4, 2); +x_80 = lean_ctor_get(x_4, 3); +x_81 = lean_ctor_get(x_4, 4); +x_82 = lean_ctor_get(x_4, 5); +lean_inc(x_82); +lean_inc(x_81); +lean_inc(x_80); +lean_inc(x_79); +lean_inc(x_78); +lean_dec(x_4); +x_83 = lean_ctor_get_uint8(x_21, 0); +x_84 = lean_ctor_get_uint8(x_21, 1); +x_85 = lean_ctor_get_uint8(x_21, 2); +x_86 = lean_ctor_get_uint8(x_21, 3); +x_87 = lean_ctor_get_uint8(x_21, 4); +x_88 = lean_ctor_get_uint8(x_21, 5); +x_89 = lean_ctor_get_uint8(x_21, 6); +x_90 = lean_ctor_get_uint8(x_21, 7); +x_91 = lean_ctor_get_uint8(x_21, 8); +x_92 = lean_ctor_get_uint8(x_21, 10); +x_93 = lean_ctor_get_uint8(x_21, 11); +if (lean_is_exclusive(x_21)) { + x_94 = x_21; +} else { + lean_dec_ref(x_21); + x_94 = lean_box(0); +} +if (lean_is_scalar(x_94)) { + x_95 = lean_alloc_ctor(0, 0, 12); +} else { + x_95 = x_94; +} +lean_ctor_set_uint8(x_95, 0, x_83); +lean_ctor_set_uint8(x_95, 1, x_84); +lean_ctor_set_uint8(x_95, 2, x_85); +lean_ctor_set_uint8(x_95, 3, x_86); +lean_ctor_set_uint8(x_95, 4, x_87); +lean_ctor_set_uint8(x_95, 5, x_88); +lean_ctor_set_uint8(x_95, 6, x_89); +lean_ctor_set_uint8(x_95, 7, x_90); +lean_ctor_set_uint8(x_95, 8, x_91); +lean_ctor_set_uint8(x_95, 9, x_2); +lean_ctor_set_uint8(x_95, 10, x_92); +lean_ctor_set_uint8(x_95, 11, x_93); +x_96 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_96, 0, x_95); +lean_ctor_set(x_96, 1, x_78); +lean_ctor_set(x_96, 2, x_79); +lean_ctor_set(x_96, 3, x_80); +lean_ctor_set(x_96, 4, x_81); +lean_ctor_set(x_96, 5, x_82); +x_97 = l_Lean_Meta_Canonicalizer_canon___closed__1; +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_96); +lean_inc(x_22); +lean_inc(x_1); +x_98 = l_List_forIn_loop___at_Lean_Meta_Canonicalizer_canon___spec__1(x_1, x_97, x_22, x_97, x_2, x_3, x_96, x_5, x_6, x_7, x_14); +if (lean_obj_tag(x_98) == 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_99, 0); +lean_inc(x_100); +lean_dec(x_99); +if (lean_obj_tag(x_100) == 0) +{ +lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_101 = lean_ctor_get(x_98, 1); +lean_inc(x_101); +lean_dec(x_98); +x_102 = lean_box(0); +x_103 = l_Lean_Meta_Canonicalizer_canon___lambda__1(x_1, x_10, x_22, x_102, x_2, x_3, x_96, x_5, x_6, x_7, x_101); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_96); +lean_dec(x_3); +x_104 = lean_ctor_get(x_103, 0); +lean_inc(x_104); +x_105 = lean_ctor_get(x_103, 1); +lean_inc(x_105); +if (lean_is_exclusive(x_103)) { + lean_ctor_release(x_103, 0); + lean_ctor_release(x_103, 1); + x_106 = x_103; +} else { + lean_dec_ref(x_103); + x_106 = lean_box(0); +} +if (lean_is_scalar(x_106)) { + x_107 = lean_alloc_ctor(0, 2, 0); +} else { + x_107 = x_106; +} +lean_ctor_set(x_107, 0, x_104); +lean_ctor_set(x_107, 1, x_105); +return x_107; +} +else +{ +lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; +lean_dec(x_96); +lean_dec(x_22); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_108 = lean_ctor_get(x_98, 1); +lean_inc(x_108); +if (lean_is_exclusive(x_98)) { + lean_ctor_release(x_98, 0); + lean_ctor_release(x_98, 1); + x_109 = x_98; +} else { + lean_dec_ref(x_98); + x_109 = lean_box(0); +} +x_110 = lean_ctor_get(x_100, 0); +lean_inc(x_110); +lean_dec(x_100); +if (lean_is_scalar(x_109)) { + x_111 = lean_alloc_ctor(0, 2, 0); +} else { + x_111 = x_109; +} +lean_ctor_set(x_111, 0, x_110); +lean_ctor_set(x_111, 1, x_108); +return x_111; +} +} +else +{ +lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; +lean_dec(x_96); +lean_dec(x_22); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_112 = lean_ctor_get(x_98, 0); +lean_inc(x_112); +x_113 = lean_ctor_get(x_98, 1); +lean_inc(x_113); +if (lean_is_exclusive(x_98)) { + lean_ctor_release(x_98, 0); + lean_ctor_release(x_98, 1); + x_114 = x_98; +} else { + lean_dec_ref(x_98); + 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(x_115, 0, x_112); +lean_ctor_set(x_115, 1, x_113); +return x_115; +} +} +} +} +else +{ +uint8_t x_116; +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_116 = !lean_is_exclusive(x_9); +if (x_116 == 0) +{ +return x_9; +} +else +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; +x_117 = lean_ctor_get(x_9, 0); +x_118 = lean_ctor_get(x_9, 1); +lean_inc(x_118); +lean_inc(x_117); +lean_dec(x_9); +x_119 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_119, 0, x_117); +lean_ctor_set(x_119, 1, x_118); +return x_119; +} +} +} +} +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Meta_Canonicalizer_canon___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_5); +lean_dec(x_5); +x_13 = l_List_forIn_loop___at_Lean_Meta_Canonicalizer_canon___spec__1(x_1, x_2, x_3, x_4, x_12, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_6); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Canonicalizer_canon___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: +{ +uint8_t x_12; lean_object* x_13; +x_12 = lean_unbox(x_5); +lean_dec(x_5); +x_13 = l_Lean_Meta_Canonicalizer_canon___lambda__1(x_1, x_2, x_3, x_4, x_12, 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_4); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Canonicalizer_canon___boxed(lean_object* x_1, lean_object* x_2, 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_Meta_Canonicalizer_canon(x_1, x_9, x_3, x_4, x_5, x_6, x_7, x_8); +return x_10; +} +} +lean_object* initialize_Lean_Util_ShareCommon(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Data_HashMap(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Basic(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_FunInfo(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Lean_Meta_Canonicalizer(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Lean_Util_ShareCommon(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Data_HashMap(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Basic(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_FunInfo(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Lean_Meta_Canonicalizer_instInhabitedExprVisited___closed__1 = _init_l_Lean_Meta_Canonicalizer_instInhabitedExprVisited___closed__1(); +lean_mark_persistent(l_Lean_Meta_Canonicalizer_instInhabitedExprVisited___closed__1); +l_Lean_Meta_Canonicalizer_instInhabitedExprVisited = _init_l_Lean_Meta_Canonicalizer_instInhabitedExprVisited(); +lean_mark_persistent(l_Lean_Meta_Canonicalizer_instInhabitedExprVisited); +l_Lean_Meta_Canonicalizer_State_keys___default___closed__1 = _init_l_Lean_Meta_Canonicalizer_State_keys___default___closed__1(); +lean_mark_persistent(l_Lean_Meta_Canonicalizer_State_keys___default___closed__1); +l_Lean_Meta_Canonicalizer_State_keys___default = _init_l_Lean_Meta_Canonicalizer_State_keys___default(); +lean_mark_persistent(l_Lean_Meta_Canonicalizer_State_keys___default); +l_Lean_Meta_Canonicalizer_State_cache___default___closed__1 = _init_l_Lean_Meta_Canonicalizer_State_cache___default___closed__1(); +lean_mark_persistent(l_Lean_Meta_Canonicalizer_State_cache___default___closed__1); +l_Lean_Meta_Canonicalizer_State_cache___default = _init_l_Lean_Meta_Canonicalizer_State_cache___default(); +lean_mark_persistent(l_Lean_Meta_Canonicalizer_State_cache___default); +l_Lean_Meta_Canonicalizer_State_keyToExprs___default = _init_l_Lean_Meta_Canonicalizer_State_keyToExprs___default(); +lean_mark_persistent(l_Lean_Meta_Canonicalizer_State_keyToExprs___default); +l_Lean_Meta_Canonicalizer_instInhabitedState___closed__1 = _init_l_Lean_Meta_Canonicalizer_instInhabitedState___closed__1(); +lean_mark_persistent(l_Lean_Meta_Canonicalizer_instInhabitedState___closed__1); +l_Lean_Meta_Canonicalizer_instInhabitedState = _init_l_Lean_Meta_Canonicalizer_instInhabitedState(); +lean_mark_persistent(l_Lean_Meta_Canonicalizer_instInhabitedState); +l_Array_mapIdxM_map___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___spec__2___closed__1 = _init_l_Array_mapIdxM_map___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___spec__2___closed__1(); +lean_mark_persistent(l_Array_mapIdxM_map___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___spec__2___closed__1); +l_Array_mapIdxM_map___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___spec__2___closed__2 = _init_l_Array_mapIdxM_map___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___spec__2___closed__2(); +lean_mark_persistent(l_Array_mapIdxM_map___at___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___spec__2___closed__2); +l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___lambda__2___closed__1 = _init_l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___lambda__2___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Canonicalizer_0__Lean_Meta_Canonicalizer_mkKey___lambda__2___closed__1); +l_Lean_Meta_Canonicalizer_canon___closed__1 = _init_l_Lean_Meta_Canonicalizer_canon___closed__1(); +lean_mark_persistent(l_Lean_Meta_Canonicalizer_canon___closed__1); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Lean/Meta/Tactic/AC/Main.c b/stage0/stdlib/Lean/Meta/Tactic/AC/Main.c index ddfe8f9f65c2..44b50bc0dfda 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/AC/Main.c +++ b/stage0/stdlib/Lean/Meta/Tactic/AC/Main.c @@ -2289,7 +2289,7 @@ static lean_object* _init_l_Lean_HashMap_find_x21___at_Lean_Meta_AC_toACExpr___s 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_HashMap_find_x21___at_Lean_Meta_AC_toACExpr___spec__13___closed__1; x_2 = l_Lean_HashMap_find_x21___at_Lean_Meta_AC_toACExpr___spec__13___closed__2; -x_3 = lean_unsigned_to_nat(184u); +x_3 = lean_unsigned_to_nat(210u); x_4 = lean_unsigned_to_nat(14u); x_5 = l_Lean_HashMap_find_x21___at_Lean_Meta_AC_toACExpr___spec__13___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/Parser/Command.c b/stage0/stdlib/Lean/Parser/Command.c index 4b78e9c3d335..354f895cc9dc 100644 --- a/stage0/stdlib/Lean/Parser/Command.c +++ b/stage0/stdlib/Lean/Parser/Command.c @@ -30,6 +30,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_initialize_parenthesizer(lean_obj static lean_object* l_Lean_Parser_Command_openRenaming___closed__6; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_openHiding_parenthesizer(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Command_mutual; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__3; static lean_object* l_Lean_Parser_Command_declaration_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Command_declId_parenthesizer___closed__10; lean_object* l_Lean_PrettyPrinter_Parenthesizer_recover_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -55,8 +56,8 @@ static lean_object* l_Lean_Parser_Command_declSig_formatter___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_end_declRange___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Command_check_formatter___closed__1; static lean_object* l_Lean_Parser_Command_deriving_parenthesizer___closed__8; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__34; LEAN_EXPORT lean_object* l_Lean_Parser_Command_extends_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__44; static lean_object* l_Lean_Parser_Command_example___closed__4; static lean_object* l_Lean_Parser_Command_namedPrio_parenthesizer___closed__11; static lean_object* l_Lean_Parser_Command_check_formatter___closed__3; @@ -68,7 +69,6 @@ static lean_object* l_Lean_Parser_Command_exit_parenthesizer___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_print(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Command_structImplicitBinder; static lean_object* l_Lean_Parser_Command_deriveInduction___closed__4; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__14; static lean_object* l_Lean_Parser_Command_printEqns_formatter___closed__1; static lean_object* l_Lean_Parser_Command_inductive___closed__21; static lean_object* l_Lean_Parser_Command_end___closed__4; @@ -109,9 +109,11 @@ 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_definition___closed__10; 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*); lean_object* l_Lean_Parser_Term_letDecl_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__29; static lean_object* l_Lean_Parser_Command_declModifiers___closed__12; static lean_object* l_Lean_Parser_Command_genInjectiveTheorems___closed__3; static lean_object* l_Lean_Parser_Command_print___closed__5; @@ -226,7 +228,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_variable_declRange___clos static lean_object* l_Lean_Parser_Command_inductive_parenthesizer___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Command_abbrev_formatter___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_private_parenthesizer(lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__30; static lean_object* l_Lean_Parser_Command_declModifiers___closed__7; static lean_object* l_Lean_Parser_Command_declModifiers___closed__14; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_private_formatter(lean_object*); @@ -240,6 +241,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Tactic_set__option(lean_obje 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___regBuiltin_Lean_Parser_Command_definition_formatter___closed__2; 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,6 +261,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_printAxioms_parenthesizer 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_openOnly_parenthesizer___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__20; 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; @@ -273,7 +276,6 @@ static lean_object* l_Lean_Parser_Command_ctor___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Command_open_declRange___closed__7; static lean_object* l_Lean_Parser_Command_printEqns_formatter___closed__6; static lean_object* l_Lean_Parser_Command_inductive_formatter___closed__10; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__27; 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*); @@ -335,13 +337,13 @@ static lean_object* l_Lean_Parser_Command_computedField_formatter___closed__2; 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; +LEAN_EXPORT lean_object* l_Lean_Parser_Command_definition_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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; static lean_object* l___regBuiltin_Lean_Parser_Command_instance_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_deriving___closed__11; static lean_object* l_Lean_Parser_Command_structFields_formatter___closed__6; @@ -351,7 +353,6 @@ static lean_object* l_Lean_Parser_Command_declValSimple_formatter___closed__2; static lean_object* l_Lean_Parser_Command_declValSimple___closed__7; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_quot_docString(lean_object*); static lean_object* l_Lean_Parser_Command_declVal___closed__6; -static lean_object* l_Lean_Parser_Command_def___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Command_declId; static lean_object* l_Lean_Parser_Command_structure_formatter___closed__17; static lean_object* l_Lean_Parser_Term_quot___closed__10; @@ -362,6 +363,8 @@ static lean_object* l_Lean_Parser_Command_whereStructInst___closed__20; static lean_object* l___regBuiltin_Lean_Parser_Command_eraseAttr_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_declId_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_synth_formatter___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__40; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__14; static lean_object* l_Lean_Parser_Command_declaration_formatter___closed__12; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_universe_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Command_theorem___closed__9; @@ -384,6 +387,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_openScoped_parenthesizer(lean_obj static lean_object* l_Lean_Parser_Tactic_set__option_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_check_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_genInjectiveTheorems_parenthesizer___closed__1; +static lean_object* l___regBuiltin_Lean_Parser_Command_definition_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_precheckedQuot_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_ctor_parenthesizer___closed__10; static lean_object* l___regBuiltin_Lean_Parser_Term_open_declRange___closed__3; @@ -398,6 +402,7 @@ static lean_object* l_Lean_Parser_Command_declaration_formatter___closed__3; static lean_object* l_Lean_Parser_Command_partial___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_definition_formatter___closed__6; static lean_object* l_Lean_Parser_Command_declId___closed__11; static lean_object* l_Lean_Parser_Command_namedPrio_parenthesizer___closed__12; static lean_object* l_Lean_Parser_Command_moduleDoc_parenthesizer___closed__6; @@ -406,12 +411,14 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_precheckedQuot_declRange___c LEAN_EXPORT lean_object* l_Lean_Parser_Command_structExplicitBinder_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_deriveInduction_parenthesizer___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Command_mutual_declRange___closed__3; +static lean_object* l_Lean_Parser_Command_definition_formatter___closed__8; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_computedFields_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Command_set__option_parenthesizer___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Term_set__option_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_quot_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_openRenamingItem___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Command_extends; +static lean_object* l_Lean_Parser_Command_definition___closed__4; static lean_object* l_Lean_Parser_Command_declaration___closed__9; static lean_object* l_Lean_Parser_Command_declaration___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Command_universe_declRange___closed__4; @@ -423,6 +430,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_precheckedQuot_declRange___c static lean_object* l___regBuiltin_Lean_Parser_Command_private_formatter___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_partial_formatter(lean_object*); lean_object* l_Lean_Parser_many_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Command_definition_formatter___closed__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; @@ -453,6 +461,7 @@ static lean_object* l_Lean_Parser_Command_structure___closed__4; static lean_object* l_Lean_Parser_Command_declVal_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Command_declSig_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_open_docString(lean_object*); +static lean_object* l_Lean_Parser_Command_definition_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_end___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_eval_parenthesizer___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_attribute(lean_object*); @@ -460,6 +469,7 @@ static lean_object* l_Lean_Parser_Command_initialize_parenthesizer___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Command_noncomputableSection_declRange___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Command_declValSimple_formatter___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Command_ctor; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__26; 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_object* l_Lean_Parser_group_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -476,6 +486,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_mutual(lean_object*) 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_definition___closed__6; static lean_object* l_Lean_Parser_Command_structure___closed__2; static lean_object* l_Lean_Parser_Command_eraseAttr___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Command_openHiding_parenthesizer___closed__1; @@ -486,7 +497,6 @@ static lean_object* l_Lean_Parser_Command_whereStructInst_parenthesizer___closed static lean_object* l_Lean_Parser_Command_check___closed__2; lean_object* l_Lean_Parser_recover(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_protected_formatter___closed__2; -static lean_object* l_Lean_Parser_Command_def_formatter___closed__8; 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; @@ -495,7 +505,6 @@ static lean_object* l_Lean_Parser_Term_set__option___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_moduleDoc_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Command_openOnly_formatter___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_classTk_formatter(lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__18; static lean_object* l_Lean_Parser_Command_eraseAttr___closed__5; static lean_object* l_Lean_Parser_Command_initializeKeyword___closed__11; static lean_object* l___regBuiltin_Lean_Parser_Command_check__failure_declRange___closed__7; @@ -531,6 +540,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Parser_Command_0__Lean_Parser_Command_ static lean_object* l_Lean_Parser_Command_declValEqns_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_structSimpleBinder_formatter___closed__1; static lean_object* l_Lean_Parser_Command_print___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__48; static lean_object* l___regBuiltin_Lean_Parser_Command_in_declRange___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Command_variable_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Command_initializeKeyword_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -551,6 +561,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_declaration_declRange___c static lean_object* l_Lean_Parser_Command_declModifiers_parenthesizer___closed__10; static lean_object* l_Lean_Parser_Command_quot___closed__12; static lean_object* l_Lean_Parser_Term_precheckedQuot_formatter___closed__4; +static lean_object* l_Lean_Parser_Command_definition_formatter___closed__1; static lean_object* l_Lean_Parser_Command_initialize_formatter___closed__10; LEAN_EXPORT lean_object* l_Lean_Parser_Command_structFields_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_quot_formatter(lean_object*); @@ -578,6 +589,7 @@ static lean_object* l_Lean_Parser_Command_moduleDoc_formatter___closed__7; static lean_object* l_Lean_Parser_Command_initializeKeyword_formatter___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Term_open_declRange___closed__1; static lean_object* l_Lean_Parser_Command_instance_formatter___closed__3; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__38; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_eval_declRange(lean_object*); static lean_object* l_Lean_Parser_Command_end___closed__5; static lean_object* l_Lean_Parser_Command_initialize_parenthesizer___closed__10; @@ -590,7 +602,6 @@ 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_computedField___closed__1; static lean_object* l_Lean_Parser_Command_nonrec___closed__6; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__46; 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; @@ -607,6 +618,7 @@ 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___regBuiltin_Lean_Parser_Command_instance_formatter___closed__2; +static lean_object* l_Lean_Parser_Command_definition___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; static lean_object* l_Lean_Parser_Command_openRenamingItem_formatter___closed__5; @@ -614,7 +626,6 @@ static lean_object* l_Lean_Parser_Command_declValEqns_formatter___closed__3; static lean_object* l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_init__quot___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Command_attribute_declRange___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__28; static lean_object* l_Lean_Parser_Command_printEqns_formatter___closed__2; static lean_object* l_Lean_Parser_Command_section___closed__1; lean_object* l_Lean_Parser_ppDedent_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -626,6 +637,7 @@ static lean_object* l_Lean_Parser_Command_mutual_formatter___closed__7; static lean_object* l_Lean_Parser_Command_section___closed__2; static lean_object* l_Lean_Parser_Command_declValSimple_formatter___closed__6; static lean_object* l_Lean_Parser_Command_structure___closed__12; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__45; static lean_object* l_Lean_Parser_Command_noncomputableSection___closed__3; static lean_object* l_Lean_Parser_Command_structImplicitBinder_parenthesizer___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Command_structureTk_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -643,7 +655,6 @@ static lean_object* l_Lean_Parser_Command_classInductive___closed__11; static lean_object* l_Lean_Parser_Command_declModifiers_formatter___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Command_computedFields_parenthesizer___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_declValSimple_parenthesizer(lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__29; static lean_object* l_Lean_Parser_Command_genInjectiveTheorems___closed__6; static lean_object* l_Lean_Parser_Command_instance___closed__1; static lean_object* l_Lean_Parser_Command_whereStructInst_formatter___closed__3; @@ -699,6 +710,7 @@ lean_object* l_Lean_Parser_incQuotDepth_formatter(lean_object*, lean_object*, le static lean_object* l___regBuiltin_Lean_Parser_Command_deriving_declRange___closed__2; static lean_object* l_Lean_Parser_Command_classInductive_parenthesizer___closed__8; extern lean_object* l_Lean_Parser_Term_matchAltsWhereDecls; +LEAN_EXPORT lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2829_(lean_object*); static lean_object* l_Lean_Parser_Command_structExplicitBinder___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Command_classInductive; static lean_object* l_Lean_Parser_Command_structImplicitBinder_formatter___closed__1; @@ -708,6 +720,7 @@ static lean_object* l_Lean_Parser_Command_optDefDeriving___closed__2; static lean_object* l_Lean_Parser_Tactic_set__option_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Command_mutual___closed__9; static lean_object* l_Lean_Parser_Command_declValSimple_parenthesizer___closed__10; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__31; static lean_object* l_Lean_Parser_Command_axiom_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_noncomputableSection___closed__10; static lean_object* l___regBuiltin_Lean_Parser_Command_check__failure_declRange___closed__5; @@ -741,6 +754,8 @@ 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_declValSimple___closed__9; static lean_object* l_Lean_Parser_Command_openOnly_parenthesizer___closed__1; +static lean_object* l___regBuiltin_Lean_Parser_Command_definition_formatter___closed__1; +static lean_object* l_Lean_Parser_Command_definition___closed__13; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_openScoped_formatter(lean_object*); static lean_object* l_Lean_Parser_Command_noncomputable___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Command_optDefDeriving; @@ -768,7 +783,6 @@ static lean_object* l_Lean_Parser_Command_nonrec_formatter___closed__1; static lean_object* l_Lean_Parser_Command_noncomputableSection___closed__4; static lean_object* l_Lean_Parser_Command_optDeriving___closed__9; LEAN_EXPORT lean_object* l_Lean_Parser_Command_mutual_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__35; static lean_object* l_Lean_Parser_Command_optDeriving_formatter___closed__2; static lean_object* l_Lean_Parser_Command_quot_parenthesizer___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Command_classTk_parenthesizer___closed__2; @@ -793,7 +807,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_structure_formatter(lean_object*, LEAN_EXPORT lean_object* l_Lean_Parser_Command_theorem; static lean_object* l_Lean_Parser_Command_namedPrio___closed__17; static lean_object* l_Lean_Parser_Command_theorem_formatter___closed__5; -static lean_object* l_Lean_Parser_Command_def_formatter___closed__6; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_quot_formatter(lean_object*); static lean_object* l_Lean_Parser_Command_instance_formatter___closed__11; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_exit_declRange(lean_object*); @@ -870,6 +883,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_variable_declRange(l static lean_object* l_Lean_Parser_Command_structFields___closed__8; lean_object* l_Lean_Parser_Term_attrInstance_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_Term_matchAltsWhereDecls_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__50; 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; @@ -885,7 +899,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_set__option_parenthesizer static lean_object* l_Lean_Parser_Command_structFields_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_exit___closed__7; static lean_object* l_Lean_Parser_Command_check__failure___closed__5; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__37; static lean_object* l_Lean_Parser_Command_initializeKeyword_formatter___closed__2; static lean_object* l_Lean_Parser_Command_declId_formatter___closed__12; static lean_object* l_Lean_Parser_Command_eoi___closed__5; @@ -910,6 +923,7 @@ 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_2851____closed__19; 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; @@ -925,6 +939,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_section_declRange___close static lean_object* l___regBuiltin_Lean_Parser_Command_deriving_declRange___closed__5; static lean_object* l_Lean_Parser_Command_check_parenthesizer___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Command_deriving_formatter___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__44; 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; @@ -951,7 +966,6 @@ static lean_object* l_Lean_Parser_Command_initialize___closed__9; 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_initFn____x40_Lean_Parser_Command___hyg_2813____closed__12; 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_printEqns_parenthesizer___closed__5; @@ -1031,8 +1045,8 @@ static lean_object* l_Lean_Parser_Command_structInstBinder___closed__14; static lean_object* l_Lean_Parser_Command_eval___closed__4; static lean_object* l_Lean_Parser_Command_printEqns_formatter___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Command_structInstBinder_parenthesizer___closed__1; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__8; static lean_object* l_Lean_Parser_Command_attribute_formatter___closed__9; +static lean_object* l_Lean_Parser_Command_definition_formatter___closed__9; LEAN_EXPORT lean_object* l_Lean_Parser_Command_genInjectiveTheorems_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_instance___closed__9; static lean_object* l___regBuiltin_Lean_Parser_Command_instance_parenthesizer___closed__2; @@ -1082,7 +1096,6 @@ lean_object* l_Lean_Parser_mkAntiquot_formatter___boxed(lean_object*, lean_objec static lean_object* l_Lean_Parser_Command_derivingClasses_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Command_openScoped___closed__4; static lean_object* l_Lean_Parser_Command_open___closed__4; -static lean_object* l_Lean_Parser_Command_def___closed__13; static lean_object* l_Lean_Parser_Command_ctor_formatter___closed__7; static lean_object* l_Lean_Parser_Command_optDeclSig_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Command_structure_parenthesizer___closed__6; @@ -1149,6 +1162,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_end_formatter(lean_object*, lean_ 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___regBuiltin_Lean_Parser_Command_initialize_declRange___closed__5; +static lean_object* l_Lean_Parser_Command_definition_parenthesizer___closed__8; 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; @@ -1171,7 +1185,6 @@ static lean_object* l_Lean_Parser_Command_check___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Command_set__option_declRange___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Command_initialize_formatter___closed__1; static lean_object* l_Lean_Parser_Command_genInjectiveTheorems_formatter___closed__1; -static lean_object* l_Lean_Parser_Command_def_formatter___closed__7; static lean_object* l_Lean_Parser_Command_structInstBinder___closed__11; static lean_object* l_Lean_Parser_Command_computedFields_formatter___closed__4; static lean_object* l_Lean_Parser_Command_addDocString_parenthesizer___closed__1; @@ -1187,9 +1200,9 @@ static lean_object* l_Lean_Parser_Command_optDefDeriving___closed__3; static lean_object* l_Lean_Parser_Command_declModifiers_formatter___closed__22; static lean_object* l_Lean_Parser_Command_axiom_parenthesizer___closed__4; lean_object* l_Lean_Parser_Term_binderIdent_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__5; static lean_object* l_Lean_Parser_Command_def_parenthesizer___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_declValEqns_formatter(lean_object*); +static lean_object* l___regBuiltin_Lean_Parser_Command_definition_parenthesizer___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_in_parenthesizer(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_attribute_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Command_whereStructInst___closed__15; @@ -1202,8 +1215,10 @@ static lean_object* l_Lean_Parser_Command_end_parenthesizer___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Command_structFields; static lean_object* l_Lean_Parser_Command_in_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_openScoped_parenthesizer___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__8; static lean_object* l_Lean_Parser_Command_genInjectiveTheorems_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Command_check__failure___closed__8; +static lean_object* l_Lean_Parser_Command_declaration_formatter___closed__14; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_variable(lean_object*); static lean_object* l_Lean_Parser_Command_classInductive_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Term_set__option___closed__8; @@ -1258,14 +1273,12 @@ static lean_object* l_Lean_Parser_Command_whereStructInst_parenthesizer___closed static lean_object* l_Lean_Parser_Command_openHiding___closed__7; 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_2813____closed__21; 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; 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_structCtor_parenthesizer___closed__4; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__23; static lean_object* l___regBuiltin_Lean_Parser_Command_export_declRange___closed__4; static lean_object* l_Lean_Parser_Command_openDecl___closed__6; static lean_object* l_Lean_Parser_Command_partial___closed__6; @@ -1306,6 +1319,7 @@ static lean_object* l_Lean_Parser_Command_openOnly___closed__8; static lean_object* l_Lean_Parser_Command_extends___closed__8; static lean_object* l_Lean_Parser_Command_structInstBinder_formatter___closed__5; static lean_object* l_Lean_Parser_Command_classInductive___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__53; 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; @@ -1353,6 +1367,7 @@ static lean_object* l_Lean_Parser_Command_structInstBinder_parenthesizer___close static lean_object* l_Lean_Parser_Command_namedPrio_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Command_namespace_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_deriveInduction_formatter___closed__4; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__46; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_addDocString_parenthesizer(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_declaration_declRange___closed__7; static lean_object* l_Lean_Parser_Command_declId___closed__3; @@ -1387,7 +1402,6 @@ static lean_object* l_Lean_Parser_Command_reduce___closed__7; static lean_object* l_Lean_Parser_Command_open_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_check__failure_formatter___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Command_derivingClasses; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__53; static lean_object* l_Lean_Parser_Tactic_open_formatter___closed__2; static lean_object* l_Lean_Parser_Command_declId___closed__7; static lean_object* l_Lean_Parser_Command_structImplicitBinder___closed__10; @@ -1414,10 +1428,12 @@ static lean_object* l_Lean_Parser_Command_openScoped___closed__6; static lean_object* l_Lean_Parser_Command_computedFields___closed__9; static lean_object* l_Lean_Parser_Command_eraseAttr_formatter___closed__2; static lean_object* l_Lean_Parser_Command_instance_formatter___closed__4; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__12; LEAN_EXPORT lean_object* l_Lean_Parser_Command_abbrev; static lean_object* l_Lean_Parser_Command_whereStructField___closed__4; static lean_object* l_Lean_Parser_Command_whereStructInst_formatter___closed__2; static lean_object* l_Lean_Parser_Command_classInductive_formatter___closed__11; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__52; static lean_object* l_Lean_Parser_Command_optDefDeriving_formatter___closed__1; static lean_object* l_Lean_Parser_Command_whereStructInst___closed__18; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_quot_declRange(lean_object*); @@ -1437,6 +1453,7 @@ static lean_object* l_Lean_Parser_Command_initializeKeyword___closed__10; static lean_object* l_Lean_Parser_Command_whereStructInst___closed__28; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_check_formatter(lean_object*); static lean_object* l_Lean_Parser_Tactic_set__option___closed__4; +static lean_object* l_Lean_Parser_Command_declaration_parenthesizer___closed__14; static lean_object* l_Lean_Parser_Command_openOnly_formatter___closed__2; static lean_object* l_Lean_Parser_Command_openRenamingItem_formatter___closed__1; static lean_object* l_Lean_Parser_Command_quot___closed__7; @@ -1453,7 +1470,6 @@ static lean_object* l___private_Lean_Parser_Command_0__Lean_Parser_Command_skipU LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_in_formatter(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_openSimple_parenthesizer___closed__1; lean_object* l_Lean_Parser_Tactic_tacticSeq_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__45; static lean_object* l___regBuiltin_Lean_Parser_Command_export_declRange___closed__3; static lean_object* l_Lean_Parser_Command_openRenamingItem___closed__7; static lean_object* l_Lean_Parser_Command_mutual_formatter___closed__10; @@ -1473,6 +1489,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_eval_parenthesizer___clos LEAN_EXPORT lean_object* l_Lean_Parser_Command_opaque; static lean_object* l_Lean_Parser_Command_declSig___closed__1; static lean_object* l_Lean_Parser_Command_structInstBinder___closed__7; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__36; static lean_object* l___regBuiltin_Lean_Parser_Command_extends_formatter___closed__1; static lean_object* l_Lean_Parser_Command_quot_formatter___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Command_noncomputable_parenthesizer___closed__1; @@ -1505,7 +1522,6 @@ static lean_object* l_Lean_Parser_Command_optDeriving_formatter___closed__3; 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_initFn____x40_Lean_Parser_Command___hyg_2813____closed__51; static lean_object* l_Lean_Parser_Command_variable_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_initialize_parenthesizer___closed__9; static lean_object* l_Lean_Parser_Command_theorem___closed__11; @@ -1519,11 +1535,11 @@ static lean_object* l_Lean_Parser_Command_instance_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_declId___closed__17; 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_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__30; 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_structSimpleBinder_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_declModifiers_parenthesizer___closed__13; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__38; 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; @@ -1531,7 +1547,6 @@ static lean_object* l_Lean_Parser_Command_whereStructInst___closed__14; static lean_object* l_Lean_Parser_Command_whereStructInst_formatter___closed__8; static lean_object* l_Lean_Parser_Command_classInductive_formatter___closed__9; static lean_object* l_Lean_Parser_Command_initialize_parenthesizer___closed__2; -LEAN_EXPORT lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813_(lean_object*); static lean_object* l_Lean_Parser_Command_declaration___closed__12; static lean_object* l_Lean_Parser_Command_computedFields___closed__8; static lean_object* l___regBuiltin_Lean_Parser_Command_mutual_declRange___closed__4; @@ -1565,6 +1580,7 @@ static lean_object* l_Lean_Parser_Command_noncomputableSection_parenthesizer___c static lean_object* l_Lean_Parser_Command_initialize_parenthesizer___closed__12; static lean_object* l___regBuiltin_Lean_Parser_Command_reduce_declRange___closed__2; static lean_object* l_Lean_Parser_Command_variable_formatter___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__9; 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; @@ -1583,7 +1599,6 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_exit_formatter(lean_ static lean_object* l___regBuiltin_Lean_Parser_Command_check__failure_declRange___closed__2; static lean_object* l_Lean_Parser_Command_reduce___closed__4; static lean_object* l_Lean_Parser_Command_structCtor___closed__1; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__40; static lean_object* l_Lean_Parser_Command_noncomputable_formatter___closed__3; static lean_object* l_Lean_Parser_Command_exit_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__10; @@ -1621,6 +1636,7 @@ static lean_object* l_Lean_Parser_Command_openScoped___closed__8; 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___regBuiltin_Lean_Parser_Command_eraseAttr_formatter___closed__1; static lean_object* l_Lean_Parser_Command_mutual___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__28; static lean_object* l_Lean_Parser_Command_eraseAttr_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Command_ctor_parenthesizer___closed__5; LEAN_EXPORT uint8_t l_Lean_Parser_Command_declId___lambda__1(uint32_t); @@ -1634,6 +1650,7 @@ static lean_object* l_Lean_Parser_Command_export_formatter___closed__2; lean_object* l_Lean_Parser_ppHardLineUnlessUngrouped_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_computedFields_formatter___closed__2; lean_object* l_Lean_Parser_sepBy1(lean_object*, lean_object*, lean_object*, uint8_t); +static lean_object* l_Lean_Parser_Command_definition_parenthesizer___closed__1; LEAN_EXPORT uint8_t l___private_Lean_Parser_Command_0__Lean_Parser_Command_skipUntilWsOrDelim___lambda__1(uint32_t); lean_object* l_Lean_Parser_notFollowedBy(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_structCtor_formatter___closed__2; @@ -1643,9 +1660,11 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_init__quot_parenthes static lean_object* l_Lean_Parser_Command_abbrev___closed__6; static lean_object* l_Lean_Parser_Command_noncomputable___closed__5; static lean_object* l_Lean_Parser_Command_attribute___closed__7; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__16; lean_object* l_Lean_Parser_takeWhileFn___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_openRenaming_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Command_printEqns___closed__5; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__39; static lean_object* l_Lean_Parser_Command_openRenaming_formatter___closed__6; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_abbrev_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Command_declVal___closed__3; @@ -1666,6 +1685,7 @@ lean_object* l_Lean_Parser_Term_typeSpec_formatter(lean_object*, lean_object*, l LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_structSimpleBinder_formatter(lean_object*); 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_definition_formatter___closed__4; static lean_object* l_Lean_Parser_Command_structFields___closed__13; static lean_object* l_Lean_Parser_Command_declVal___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Command_addDocString; @@ -1721,6 +1741,7 @@ static lean_object* l_Lean_Parser_Command_derivingClasses_parenthesizer___closed static lean_object* l___regBuiltin_Lean_Parser_Command_printEqns_declRange___closed__5; 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___regBuiltin_Lean_Parser_Command_definition_formatter(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Command_open_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_set__option___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Command_noncomputableSection_declRange___closed__7; @@ -1728,6 +1749,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_noncomputable_parent static lean_object* l_Lean_Parser_Command_inductive_parenthesizer___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Command_open_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_structure_formatter___closed__10; +static lean_object* l_Lean_Parser_Command_definition___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Command_initializeKeyword_formatter___closed__2; static lean_object* l_Lean_Parser_Command_declModifiers_formatter___closed__16; static lean_object* l_Lean_Parser_Command_inductive_parenthesizer___closed__6; @@ -1760,6 +1782,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_protected_parenthesizer__ 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_declaration___closed__18; 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; @@ -1767,6 +1790,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_attribute_declRange___clo static lean_object* l___regBuiltin_Lean_Parser_Command_check_declRange___closed__7; static lean_object* l_Lean_Parser_Command_noncomputableSection_formatter___closed__1; static lean_object* l_Lean_Parser_Command_inductive___closed__10; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__2; static lean_object* l_Lean_Parser_Command_printAxioms_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_structure_formatter___closed__18; static lean_object* l_Lean_Parser_Command_declSig_parenthesizer___closed__3; @@ -1779,7 +1803,6 @@ 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; 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_initFn____x40_Lean_Parser_Command___hyg_2813____closed__47; 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*); static lean_object* l___regBuiltin_Lean_Parser_Command_set__option_declRange___closed__5; @@ -1833,6 +1856,8 @@ static lean_object* l_Lean_Parser_Command_synth_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_declaration_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_genInjectiveTheorems_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_eoi_parenthesizer___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__5; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__25; 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; @@ -1840,6 +1865,7 @@ 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_inductive_parenthesizer___closed__13; static lean_object* l_Lean_Parser_Command_quot___closed__3; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__54; static lean_object* l_Lean_Parser_Command_partial___closed__7; static lean_object* l_Lean_Parser_Command_structCtor___closed__7; static lean_object* l_Lean_Parser_Command_initialize___closed__2; @@ -1863,7 +1889,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_2813____closed__16; static lean_object* l___private_Lean_Parser_Command_0__Lean_Parser_Command_skipUntilWsOrDelim___closed__1; static lean_object* l_Lean_Parser_Command_synth___closed__9; static lean_object* l_Lean_Parser_Command_optDefDeriving___closed__9; @@ -1915,9 +1940,9 @@ lean_object* l_Lean_PrettyPrinter_Formatter_checkColGt_formatter___boxed(lean_ob static lean_object* l_Lean_Parser_Command_export___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Command_derivingClasses_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_genInjectiveTheorems_declRange___closed__4; +static lean_object* l_Lean_Parser_Command_definition___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Tactic_open_declRange___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Command_private_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__19; LEAN_EXPORT lean_object* l_Lean_Parser_Command_deriveInduction_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_initialize___closed__5; static lean_object* l_Lean_Parser_Tactic_open_parenthesizer___closed__3; @@ -1943,13 +1968,16 @@ static lean_object* l_Lean_Parser_Command_initialize___closed__11; static lean_object* l_Lean_Parser_Command_noncomputable_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_genInjectiveTheorems_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_optDeriving___closed__4; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__42; static lean_object* l_Lean_Parser_Command_namespace_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Command_namedPrio_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_universe_declRange___closed__7; +static lean_object* l_Lean_Parser_Command_definition_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_structure_parenthesizer___closed__10; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__33; static lean_object* l_Lean_Parser_Command_opaque_parenthesizer___closed__6; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__18; static lean_object* l_Lean_Parser_Term_set__option_formatter___closed__1; +LEAN_EXPORT lean_object* l_Lean_Parser_Command_definition; static lean_object* l_Lean_Parser_Command_declModifiers_parenthesizer___closed__12; 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; @@ -1975,16 +2003,17 @@ static lean_object* l_Lean_Parser_Command_structExplicitBinder___closed__9; static lean_object* l___regBuiltin_Lean_Parser_Command_deriving_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_end_declRange___closed__3; static lean_object* l_Lean_Parser_Command_structureTk___closed__7; +static lean_object* l_Lean_Parser_Command_definition_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Command_ctor___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Command_axiom_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_axiom___closed__9; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__43; static lean_object* l_Lean_Parser_Command_initialize_formatter___closed__1; static lean_object* l_Lean_Parser_Command_optDefDeriving_parenthesizer___closed__3; lean_object* l_Lean_Parser_Term_ident_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_quot___closed__4; static lean_object* l_Lean_Parser_Command_initialize___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Command_optNamedPrio; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__32; static lean_object* l_Lean_Parser_Command_structImplicitBinder___closed__9; static lean_object* l_Lean_Parser_Command_synth_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_exit_parenthesizer___closed__2; @@ -1998,11 +2027,13 @@ 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; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_eoi_formatter(lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__35; static lean_object* l_Lean_Parser_Command_export___closed__8; static lean_object* l___regBuiltin_Lean_Parser_Command_print_formatter___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Command_export_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_ctor_parenthesizer___closed__9; static lean_object* l_Lean_Parser_Command_declVal_parenthesizer___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__24; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_open_declRange(lean_object*); static lean_object* l_Lean_Parser_Term_precheckedQuot_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_declId_formatter___closed__7; @@ -2032,6 +2063,7 @@ static lean_object* l_Lean_Parser_Command_structSimpleBinder___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_namedPrio_formatter(lean_object*); static lean_object* l_Lean_Parser_Command_opaque___closed__8; static lean_object* l_Lean_Parser_Command_inductive___closed__12; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Command_computedFields_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_inductive___closed__18; static lean_object* l_Lean_Parser_Command_noncomputable_formatter___closed__2; @@ -2039,9 +2071,9 @@ static lean_object* l_Lean_Parser_Command_declValSimple___closed__3; static lean_object* l_Lean_Parser_Command_openHiding_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Command_openOnly_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_List_elem___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__41; static lean_object* l_Lean_Parser_Command_addDocString___closed__9; static lean_object* l_Lean_Parser_Command_initializeKeyword___closed__6; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__50; static lean_object* l_Lean_Parser_Command_theorem_formatter___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Command_noncomputable_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_precheckedQuot_declRange___closed__1; @@ -2058,12 +2090,14 @@ 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_structImplicitBinder___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__22; static lean_object* l_Lean_Parser_Command_init__quot___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Command_unsafe; static lean_object* l___regBuiltin_Lean_Parser_Command_set__option_declRange___closed__3; static lean_object* l_Lean_Parser_Command_initializeKeyword___closed__7; static lean_object* l_Lean_Parser_Command_openRenamingItem___closed__10; static lean_object* l_Lean_Parser_Command_optDeclSig_formatter___closed__9; +static lean_object* l_Lean_Parser_Command_definition___closed__8; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_declValSimple_formatter(lean_object*); static lean_object* l_Lean_Parser_Command_computedFields___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_structFields_parenthesizer(lean_object*); @@ -2119,6 +2153,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_nonrec_parenthesizer lean_object* l_Lean_Parser_andthen(lean_object*, lean_object*); 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_definition_parenthesizer___closed__5; 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; @@ -2131,6 +2166,7 @@ static lean_object* l_Lean_Parser_Command_moduleDoc___closed__1; static lean_object* l_Lean_Parser_Command_structInstBinder_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_whereStructInst___closed__12; static lean_object* l_Lean_Parser_Command_structure_formatter___closed__16; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__15; static lean_object* l_Lean_Parser_Command_opaque___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Command_noncomputable; static lean_object* l___regBuiltin_Lean_Parser_Term_open_declRange___closed__4; @@ -2146,6 +2182,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_universe; static lean_object* l___regBuiltin_Lean_Parser_Command_exit_declRange___closed__7; static lean_object* l_Lean_Parser_Command_print_parenthesizer___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Command_check__failure_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__7; static lean_object* l_Lean_Parser_Command_declaration___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Command_declValSimple_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_structExplicitBinder_formatter___closed__3; @@ -2177,7 +2214,6 @@ static lean_object* l_Lean_Parser_Command_structExplicitBinder_formatter___close static lean_object* l_Lean_Parser_Command_check___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Command_ctor_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_partial_formatter___closed__1; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__36; static lean_object* l___regBuiltin_Lean_Parser_Command_set__option_declRange___closed__2; static lean_object* l_Lean_Parser_Command_private___closed__3; lean_object* l_Lean_Parser_ppSpace_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -2219,7 +2255,6 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_mutual_declRange(lea LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_deriveInduction_formatter(lean_object*); static lean_object* l_Lean_Parser_Command_whereStructInst___closed__3; static lean_object* l_Lean_Parser_Command_export___closed__6; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__31; extern lean_object* l_Lean_PrettyPrinter_Formatter_formatterAliasesRef; static lean_object* l___regBuiltin_Lean_Parser_Tactic_open_declRange___closed__5; static lean_object* l_Lean_Parser_Command_unsafe_formatter___closed__3; @@ -2227,10 +2262,10 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_deriveInduction_declRange static lean_object* l___regBuiltin_Lean_Parser_Command_export_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Tactic_open_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_ctor_formatter___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__21; static lean_object* l___regBuiltin_Lean_Parser_Command_reduce_declRange___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Term_quot_parenthesizer___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Command_theorem_formatter___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__20; LEAN_EXPORT lean_object* l_Lean_Parser_Command_in_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_structure___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_openScoped_parenthesizer(lean_object*); @@ -2250,10 +2285,8 @@ static lean_object* l_Lean_Parser_Command_reduce___closed__1; static lean_object* l_Lean_Parser_Command_computedFields_parenthesizer___closed__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_initFn____x40_Lean_Parser_Command___hyg_2813____closed__41; 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_2813____closed__39; 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; @@ -2279,10 +2312,8 @@ static lean_object* l_Lean_Parser_Command_optDeriving___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_moduleDoc_declRange___closed__3; static lean_object* l_Lean_Parser_Command_whereStructInst___closed__6; static lean_object* l_Lean_Parser_Command_eraseAttr_formatter___closed__3; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__9; static lean_object* l_Lean_Parser_Command_openHiding_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Command_check__failure_formatter___closed__3; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__49; lean_object* l_Lean_Parser_many(lean_object*); 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*); @@ -2330,8 +2361,8 @@ static lean_object* l_Lean_Parser_Command_check__failure___closed__4; 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; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__27; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_import_formatter(lean_object*); -static lean_object* l_Lean_Parser_Command_def_formatter___closed__9; static lean_object* l_Lean_Parser_Command_structFields_formatter___closed__7; static lean_object* l_Lean_Parser_Command_unsafe___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Command_variable; @@ -2340,7 +2371,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_nonrec; static lean_object* l_Lean_Parser_Command_declValEqns___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_initialize_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Command_export_parenthesizer___closed__2; -static lean_object* l_Lean_Parser_Command_def___closed__8; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_def_formatter(lean_object*); static lean_object* l_Lean_Parser_Command_declValSimple_parenthesizer___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Command_whereStructField_formatter___closed__2; @@ -2380,11 +2410,9 @@ static lean_object* l_Lean_Parser_Command_optDeclSig___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Command_theorem_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_structureTk___closed__5; static lean_object* l_Lean_Parser_Tactic_open_formatter___closed__1; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__6; lean_object* l_Lean_Parser_Termination_suffix_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_openScoped___closed__3; static lean_object* l_Lean_Parser_Command_moduleDoc_parenthesizer___closed__2; -LEAN_EXPORT lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2791_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Command_whereStructField_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_printAxioms_declRange___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Term_open_declRange___closed__2; @@ -2428,6 +2456,7 @@ static lean_object* l_Lean_Parser_Command_private___closed__1; extern lean_object* l_Lean_Parser_ident; LEAN_EXPORT lean_object* l_Lean_Parser_Command_eoi_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_opaque_parenthesizer___closed__3; +static lean_object* l_Lean_Parser_Command_definition___closed__11; 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_Command_derivingClasses_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_import___closed__4; @@ -2441,7 +2470,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_printAxioms_formatter___c 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*); static lean_object* l_Lean_Parser_Command_namedPrio_parenthesizer___closed__9; -static lean_object* l_Lean_Parser_Command_def___closed__9; static lean_object* l_Lean_Parser_Command_structImplicitBinder_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_declId_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_classTk___closed__5; @@ -2474,7 +2502,6 @@ static lean_object* l_Lean_Parser_Command_eval_formatter___closed__4; static lean_object* l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_initialize_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Command_structFields_parenthesizer___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__24; static lean_object* l_Lean_Parser_Command_declSig_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Command_unsafe_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_genInjectiveTheorems___closed__1; @@ -2497,7 +2524,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_namespace_declRange___clo static lean_object* l___regBuiltin_Lean_Parser_Command_deriveInduction_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_deriveInduction_declRange___closed__4; static lean_object* l_Lean_Parser_Command_openRenamingItem_parenthesizer___closed__4; -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*); @@ -2530,8 +2556,9 @@ static lean_object* l_Lean_Parser_Command_structure_formatter___closed__13; static lean_object* l_Lean_Parser_Command_openDecl___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_printAxioms_declRange___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Command_namespace_declRange___closed__4; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__48; +static lean_object* l_Lean_Parser_Command_definition___closed__12; static lean_object* l___regBuiltin_Lean_Parser_Command_theorem_formatter___closed__1; +static lean_object* l_Lean_Parser_Command_definition___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Command_structSimpleBinder_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_init__quot_declRange___closed__5; static lean_object* l_Lean_Parser_Command_mutual_parenthesizer___closed__8; @@ -2550,7 +2577,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_structureTk_formatter(lean_object static lean_object* l_Lean_Parser_Command_structInstBinder_formatter___closed__1; static lean_object* l_Lean_Parser_Command_structure_formatter___closed__8; static lean_object* l_Lean_Parser_Command_deriveInduction_parenthesizer___closed__1; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__17; static lean_object* l_Lean_Parser_Command_declValEqns_formatter___closed__1; static lean_object* l_Lean_Parser_Command_derivingClasses_formatter___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Command_openDecl_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -2564,6 +2590,7 @@ static lean_object* l_Lean_Parser_Command_openDecl___closed__8; lean_object* l_Lean_Parser_withCache(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_mutual_formatter___closed__4; static lean_object* l_Lean_Parser_Command_structure___closed__14; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_openRenaming_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Command_optionValue_formatter___closed__4; static lean_object* l_Lean_Parser_Command_theorem_formatter___closed__4; @@ -2666,6 +2693,7 @@ lean_object* l_Lean_Parser_Term_binderDefault_formatter(lean_object*, lean_objec static lean_object* l_Lean_Parser_Command_structure_formatter___closed__19; static lean_object* l___regBuiltin_Lean_Parser_Command_nonrec_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_optDeclSig_formatter___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__11; static lean_object* l_Lean_Parser_Command_whereStructField_formatter___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Command_print_declRange___closed__4; static lean_object* l_Lean_Parser_Term_quot_formatter___closed__5; @@ -2686,6 +2714,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_eval_declRange___closed__ static lean_object* l_Lean_Parser_Command_computedFields_formatter___closed__7; static lean_object* l_Lean_Parser_Command_eoi_formatter___closed__5; static lean_object* l_Lean_Parser_Command_openHiding_parenthesizer___closed__4; +static lean_object* l_Lean_Parser_Command_definition___closed__7; lean_object* l_Lean_Parser_Term_doSeq_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_registerAlias(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_attribute_parenthesizer___closed__8; @@ -2711,7 +2740,6 @@ static lean_object* l_Lean_Parser_Command_structExplicitBinder_formatter___close static lean_object* l_Lean_Parser_Command_instance___closed__5; static lean_object* l_Lean_Parser_Command_quot_formatter___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_declId_formatter(lean_object*); -static lean_object* l_Lean_Parser_Command_def_parenthesizer___closed__9; static lean_object* l___regBuiltin_Lean_Parser_Command_declaration_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_attribute_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_optDeriving_parenthesizer___closed__6; @@ -2739,6 +2767,7 @@ static lean_object* l_Lean_Parser_Command_openScoped___closed__2; static lean_object* l_Lean_Parser_Command_initialize_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Tactic_set__option_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_quot_formatter___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__42; static lean_object* l_Lean_Parser_Command_quot_parenthesizer___closed__3; lean_object* l_Lean_Parser_many1(lean_object*); static lean_object* l_Lean_Parser_Command_example___closed__3; @@ -2784,6 +2813,7 @@ static lean_object* l_Lean_Parser_Command_namedPrio_formatter___closed__12; static lean_object* l_Lean_Parser_Command_printAxioms_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_addDocString_formatter___closed__1; static lean_object* l_Lean_Parser_Command_exit___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__6; static lean_object* l_Lean_Parser_Command_quot___closed__9; static lean_object* l___regBuiltin_Lean_Parser_Command_declaration_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_deriving_parenthesizer___closed__3; @@ -2843,17 +2873,16 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_exit_parenthesizer___clos static lean_object* l_Lean_Parser_Command_namedPrio_formatter___closed__8; static lean_object* l_Lean_Parser_Command_moduleDoc_parenthesizer___closed__8; static lean_object* l_Lean_Parser_Command_openHiding_parenthesizer___closed__1; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__22; LEAN_EXPORT lean_object* l_Lean_Parser_Command_instance_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_ctor___closed__8; static lean_object* l_Lean_Parser_Command_open___closed__9; static lean_object* l___regBuiltin_Lean_Parser_Command_initialize_declRange___closed__1; static lean_object* l_Lean_Parser_Command_export___closed__9; static lean_object* l_Lean_Parser_Command_opaque_parenthesizer___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__26; 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_definition___closed__9; 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; @@ -2870,6 +2899,7 @@ static lean_object* l_Lean_Parser_Command_declId_formatter___closed__1; lean_object* l_Lean_Parser_Term_leftArrow_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_initialize_declRange___closed__4; static lean_object* l_Lean_Parser_Command_declValSimple___closed__10; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__37; static lean_object* l_Lean_Parser_Command_deriving___closed__9; static lean_object* l_Lean_Parser_Command_opaque_formatter___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Command_declModifiers(uint8_t); @@ -2887,10 +2917,10 @@ static lean_object* l_Lean_Parser_Command_structInstBinder_parenthesizer___close static lean_object* l_Lean_Parser_Command_declaration_formatter___closed__10; static lean_object* l___regBuiltin_Lean_Parser_Command_print_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_structure___closed__7; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__47; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Tactic_open_formatter(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_optDeriving_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_eval___closed__8; -static lean_object* l_Lean_Parser_Command_def___closed__10; static lean_object* l___regBuiltin_Lean_Parser_Tactic_open___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_initialize_formatter___closed__2; static lean_object* l_Lean_Parser_Command_inductive___closed__2; @@ -2912,7 +2942,6 @@ 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_print_formatter___closed__6; static lean_object* l_Lean_Parser_Command_openHiding_parenthesizer___closed__6; -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___regBuiltin_Lean_Parser_Command_deriveInduction_declRange___closed__5; static lean_object* l_Lean_Parser_Command_derivingClasses_formatter___closed__5; @@ -2940,7 +2969,6 @@ lean_object* l_Lean_Parser_Term_whereDecls_parenthesizer(lean_object*, lean_obje static lean_object* l_Lean_Parser_Command_printEqns___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_axiom_parenthesizer(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Term_quot_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_def_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_theorem_formatter___closed__6; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_structureTk_parenthesizer(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_whereStructInst_parenthesizer(lean_object*); @@ -2975,9 +3003,7 @@ static lean_object* l_Lean_Parser_Command_extends_parenthesizer___closed__2; 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; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Command_openRenaming_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__34; static lean_object* l_Lean_Parser_Command_structSimpleBinder___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Command_printEqns_declRange___closed__6; static lean_object* l_Lean_Parser_Command_print___closed__6; @@ -2985,9 +3011,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_declaration_formatter(lean_object static lean_object* l_Lean_Parser_Command_declaration___closed__7; static lean_object* l_Lean_Parser_Command_mutual___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Command_optDeclSig_parenthesizer___closed__2; -static lean_object* l_Lean_Parser_Command_def___closed__11; static lean_object* l_Lean_Parser_Command_openDecl_parenthesizer___closed__1; -static lean_object* l_Lean_Parser_Command_def___closed__12; static lean_object* l_Lean_Parser_Command_init__quot_parenthesizer___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Command_variable_declRange___closed__7; static lean_object* l_Lean_Parser_Command_instance_parenthesizer___closed__9; @@ -2995,6 +3019,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_genInjectiveTheorems_decl lean_object* l_Lean_Parser_Term_binderDefault_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_computedFields___closed__7; static lean_object* l_Lean_Parser_Command_inductive___closed__1; +static lean_object* l_Lean_Parser_Command_definition_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Term_quot___closed__1; static lean_object* l_Lean_Parser_Command_deriveInduction___closed__1; static lean_object* l_Lean_Parser_Command_extends_formatter___closed__4; @@ -3028,8 +3053,8 @@ static lean_object* l_Lean_Parser_Term_quot___closed__9; static lean_object* l_Lean_Parser_Command_exit_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_inductive_parenthesizer___closed__9; static lean_object* l_Lean_Parser_Command_computedFields_parenthesizer___closed__7; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__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_genInjectiveTheorems___closed__4; @@ -3045,6 +3070,7 @@ static lean_object* l_Lean_Parser_Command_whereStructInst___closed__22; static lean_object* l_Lean_Parser_Command_abbrev___closed__1; static lean_object* l_Lean_Parser_Command_declId___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Command_section_formatter___closed__2; +static lean_object* l_Lean_Parser_Command_definition_formatter___closed__2; static lean_object* l_Lean_Parser_Command_namedPrio_parenthesizer___closed__8; LEAN_EXPORT lean_object* l_Lean_Parser_Command_set__option_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_mutual_parenthesizer___closed__6; @@ -3102,7 +3128,9 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_addDocString_formatter(lean_objec static lean_object* l_Lean_Parser_Command_abbrev_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Command_eraseAttr___closed__6; static lean_object* l_Lean_Parser_Command_set__option___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__23; static lean_object* l_Lean_Parser_Command_inductive___closed__9; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__49; static lean_object* l_Lean_Parser_Command_moduleDoc_formatter___closed__6; static lean_object* l_Lean_Parser_Term_open_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_openHiding_parenthesizer___closed__3; @@ -3120,7 +3148,6 @@ static lean_object* l_Lean_Parser_Command_noncomputableSection_formatter___close static lean_object* l_Lean_Parser_Command_declId_formatter___closed__5; static lean_object* l_Lean_Parser_Command_attribute_formatter___closed__5; static lean_object* l_Lean_Parser_Command_openRenamingItem___closed__1; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__4; static lean_object* l_Lean_Parser_Command_import___closed__6; static lean_object* l_Lean_Parser_Command_abbrev_parenthesizer___closed__8; static lean_object* l_Lean_Parser_Command_noncomputable___closed__8; @@ -3136,13 +3163,16 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_init__quot_formatter LEAN_EXPORT lean_object* l_Lean_Parser_Command_universe_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_incQuotDepth_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_in_declRange___closed__4; +static lean_object* l_Lean_Parser_Command_definition_parenthesizer___closed__9; 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_export_formatter___closed__3; static lean_object* l_Lean_Parser_Command_abbrev___closed__2; +static lean_object* l_Lean_Parser_Command_definition_formatter___closed__5; static lean_object* l_Lean_Parser_Command_whereStructInst___closed__8; static lean_object* l_Lean_Parser_Command_whereStructField_formatter___closed__2; static lean_object* l_Lean_Parser_Command_unsafe_formatter___closed__1; +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_definition_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Command_end_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Command_structureTk___closed__1; static lean_object* l_Lean_Parser_Command_ctor_formatter___closed__2; @@ -3167,12 +3197,10 @@ lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_o static lean_object* l_Lean_Parser_Command_addDocString___closed__1; lean_object* l_Lean_Parser_strLit_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_universe___closed__5; -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_initFn____x40_Lean_Parser_Command___hyg_2813____closed__33; static lean_object* l_Lean_Parser_Command_structCtor_formatter___closed__7; lean_object* l_instBEq___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_export_formatter___closed__5; @@ -3182,7 +3210,6 @@ static lean_object* l_Lean_Parser_Command_inductive_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_ctor___closed__1; static lean_object* l_Lean_Parser_Command_declModifiers___closed__6; static lean_object* l_Lean_Parser_Command_eraseAttr___closed__8; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__2; static lean_object* l_Lean_Parser_Command_optionValue___closed__4; static lean_object* l_Lean_Parser_Command_mutual___closed__10; LEAN_EXPORT lean_object* l_Lean_Parser_Command_declModifiers_formatter(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -3193,7 +3220,6 @@ static lean_object* l_Lean_Parser_Command_inductive_formatter___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Command_example_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_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_2813____closed__11; static lean_object* l_Lean_Parser_Command_namedPrio___closed__12; static lean_object* l_Lean_Parser_Command_instance___closed__4; static lean_object* l_Lean_Parser_Command_declId___closed__18; @@ -3204,6 +3230,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_in_declRange___closed__2; static lean_object* l_Lean_Parser_Command_mutual_formatter___closed__9; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_open_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Term_open_formatter___closed__4; +static lean_object* l_Lean_Parser_Command_definition_formatter___closed__3; static lean_object* l_Lean_Parser_Command_universe_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Command_end_formatter___closed__1; extern lean_object* l_Lean_Parser_Term_optType; @@ -3217,7 +3244,6 @@ static lean_object* l_Lean_Parser_Command_declValSimple_parenthesizer___closed__ static lean_object* l_Lean_Parser_Command_declModifiers_parenthesizer___closed__24; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_abbrev_formatter(lean_object*); static lean_object* l_Lean_Parser_Command_optDeclSig_formatter___closed__1; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__10; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_instance_formatter(lean_object*); static lean_object* l_Lean_Parser_Term_set__option_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_structCtor_formatter___closed__4; @@ -3349,7 +3375,6 @@ static lean_object* l_Lean_Parser_Command_optDeriving_parenthesizer___closed__4; 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; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__3; static lean_object* l_Lean_Parser_Command_exit_formatter___closed__2; static lean_object* l_Lean_Parser_Command_in_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_abbrev_formatter___closed__8; @@ -3360,11 +3385,11 @@ static lean_object* l_Lean_Parser_Command_quot___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Command_exit_declRange___closed__1; static lean_object* l_Lean_Parser_Command_classInductive___closed__16; static lean_object* l_Lean_Parser_Command_computedFields_parenthesizer___closed__4; +LEAN_EXPORT lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851_(lean_object*); 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_declaration_parenthesizer___closed__13; static lean_object* l_Lean_Parser_Command_openOnly___closed__6; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__43; static lean_object* l_Lean_Parser_Command_structInstBinder_formatter___closed__8; static lean_object* l_Lean_Parser_Command_print_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Command_import_formatter___closed__2; @@ -3379,7 +3404,6 @@ 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_derivingClasses_formatter___closed__1; static lean_object* l_Lean_Parser_Command_declaration___closed__17; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__1; static lean_object* l_Lean_Parser_Command_openRenaming___closed__10; static lean_object* l_Lean_Parser_Command_set__option___closed__1; static lean_object* l_Lean_Parser_Term_quot_formatter___closed__9; @@ -3402,7 +3426,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_abbrev_parenthesizer(lean_object* static lean_object* l___regBuiltin_Lean_Parser_Command_openOnly_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_declValEqns_parenthesizer___closed__1; lean_object* l_Lean_Parser_Term_bracketedBinder(uint8_t); -static lean_object* l_Lean_Parser_Command_def_formatter___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Command_optDeriving; static lean_object* l___regBuiltin_Lean_Parser_Command_axiom_formatter___closed__1; static lean_object* l_Lean_Parser_Command_declId_parenthesizer___closed__4; @@ -3417,10 +3440,12 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_optDeclSig_formatter___cl static lean_object* l_Lean_Parser_Command_section_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_instance___closed__7; static lean_object* l_Lean_Parser_Command_attribute_formatter___closed__7; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__51; static lean_object* l_Lean_Parser_Command_classTk___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_moduleDoc_declRange___closed__6; static lean_object* l_Lean_Parser_Command_derivingClasses___closed__3; static lean_object* l_Lean_Parser_Command_whereStructField___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__10; static lean_object* l_Lean_Parser_Command_variable_formatter___closed__2; static lean_object* l_Lean_Parser_Command_universe_formatter___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_noncomputableSection(lean_object*); @@ -3432,7 +3457,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_addDocString_parenthesize static lean_object* l_Lean_Parser_Command_attribute___closed__14; static lean_object* l_Lean_Parser_Command_classInductive_parenthesizer___closed__9; static lean_object* l_Lean_Parser_Command_addDocString_formatter___closed__4; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__15; static lean_object* l_Lean_Parser_Command_structSimpleBinder___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Command_ctor_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_declModifiers_formatter___closed__25; @@ -3442,11 +3466,13 @@ static lean_object* l_Lean_Parser_Command_initialize___closed__8; static lean_object* l_Lean_Parser_Command_declaration___closed__4; static lean_object* l_Lean_Parser_Term_quot_formatter___closed__1; static lean_object* l_Lean_Parser_Command_openRenaming_parenthesizer___closed__7; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__32; static lean_object* l_Lean_Parser_Command_whereStructField_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_declModifiers_parenthesizer___closed__25; LEAN_EXPORT lean_object* l_Lean_Parser_Command_structureTk; static lean_object* l_Lean_Parser_Command_mutual___closed__12; static lean_object* l_Lean_Parser_Command_quot_parenthesizer___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__17; static lean_object* l_Lean_Parser_Command_structImplicitBinder_parenthesizer___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Command_eoi_formatter___closed__1; static lean_object* l_Lean_Parser_Term_quot_parenthesizer___closed__9; @@ -3462,10 +3488,10 @@ static lean_object* l_Lean_Parser_Command_abbrev_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Command_structure___closed__6; static lean_object* l_Lean_Parser_Command_inductive_formatter___closed__1; static lean_object* l_Lean_Parser_Command_inductive___closed__20; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__52; LEAN_EXPORT lean_object* l_Lean_Parser_Command_namespace; static lean_object* l___regBuiltin_Lean_Parser_Command_print_formatter___closed__1; static lean_object* l_Lean_Parser_Command_mutual_parenthesizer___closed__1; +LEAN_EXPORT lean_object* l_Lean_Parser_Command_definition_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_computedField_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Command_structSimpleBinder___closed__9; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_openSimple_formatter(lean_object*); @@ -3504,11 +3530,9 @@ static lean_object* l_Lean_Parser_Command_declModifiers_formatter___closed__12; static lean_object* l_Lean_Parser_Command_structure___closed__10; static lean_object* l_Lean_Parser_Command_structExplicitBinder_formatter___closed__14; static lean_object* l_Lean_Parser_Command_namespace_parenthesizer___closed__2; -static lean_object* l_Lean_Parser_Command_def_parenthesizer___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Command_section_declRange___closed__5; static lean_object* l_Lean_Parser_Command_namespace___closed__5; static lean_object* l_Lean_Parser_Command_openRenamingItem_parenthesizer___closed__5; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__25; static lean_object* l___regBuiltin_Lean_Parser_Command_noncomputable_formatter___closed__1; static lean_object* l_Lean_Parser_Command_inductive_parenthesizer___closed__10; LEAN_EXPORT lean_object* l_Lean_Parser_Command_visibility; @@ -3535,9 +3559,7 @@ static lean_object* l_Lean_Parser_Term_precheckedQuot_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_unsafe_parenthesizer___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Command_open_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_structure_parenthesizer___closed__4; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__54; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_in(lean_object*); -static lean_object* l_Lean_Parser_Command_def_parenthesizer___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Command_printAxioms_declRange___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Tactic_set__option_declRange___closed__5; static lean_object* l_Lean_Parser_Command_optDeclSig_formatter___closed__3; @@ -3582,7 +3604,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_set__option_parenthesizer(lean_objec static lean_object* l_Lean_Parser_Command_instance_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_deriveInduction_declRange___closed__1; static lean_object* l_Lean_Parser_Command_axiom_formatter___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__13; static lean_object* l_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Command_printAxioms_formatter___closed__4; static lean_object* l_Lean_Parser_Command_whereStructInst___closed__9; @@ -3593,6 +3614,7 @@ static lean_object* l_Lean_Parser_Command_genInjectiveTheorems_formatter___close 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_private___closed__6; +static lean_object* l_Lean_Parser_Command_definition_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Command_exit___closed__2; static lean_object* l_Lean_Parser_Command_structInstBinder___closed__1; lean_object* l_Lean_Parser_termParser(lean_object*); @@ -8255,39 +8277,39 @@ x_1 = l_Lean_Parser_Command_optDefDeriving___closed__11; return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_def___closed__1() { +static lean_object* _init_l_Lean_Parser_Command_definition___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("def", 3); +x_1 = lean_mk_string_from_bytes("definition", 10); return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_def___closed__2() { +static lean_object* _init_l_Lean_Parser_Command_definition___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_def___closed__1; +x_4 = l_Lean_Parser_Command_definition___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_Command_def___closed__3() { +static lean_object* _init_l_Lean_Parser_Command_definition___closed__3() { _start: { lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; -x_1 = l_Lean_Parser_Command_def___closed__1; -x_2 = l_Lean_Parser_Command_def___closed__2; +x_1 = l_Lean_Parser_Command_definition___closed__1; +x_2 = l_Lean_Parser_Command_definition___closed__2; x_3 = 1; x_4 = 0; x_5 = l_Lean_Parser_mkAntiquot(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Parser_Command_def___closed__4() { +static lean_object* _init_l_Lean_Parser_Command_definition___closed__4() { _start: { lean_object* x_1; @@ -8295,16 +8317,16 @@ x_1 = lean_mk_string_from_bytes("def ", 4); return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_def___closed__5() { +static lean_object* _init_l_Lean_Parser_Command_definition___closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_def___closed__4; +x_1 = l_Lean_Parser_Command_definition___closed__4; x_2 = l_Lean_Parser_symbol(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_def___closed__6() { +static lean_object* _init_l_Lean_Parser_Command_definition___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -8314,7 +8336,7 @@ x_3 = l_Lean_Parser_recover(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_def___closed__7() { +static lean_object* _init_l_Lean_Parser_Command_definition___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -8324,63 +8346,134 @@ x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_def___closed__8() { +static lean_object* _init_l_Lean_Parser_Command_definition___closed__8() { _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_def___closed__7; +x_2 = l_Lean_Parser_Command_definition___closed__7; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_def___closed__9() { +static lean_object* _init_l_Lean_Parser_Command_definition___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_def___closed__6; -x_2 = l_Lean_Parser_Command_def___closed__8; +x_1 = l_Lean_Parser_Command_definition___closed__6; +x_2 = l_Lean_Parser_Command_definition___closed__8; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_def___closed__10() { +static lean_object* _init_l_Lean_Parser_Command_definition___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_def___closed__5; -x_2 = l_Lean_Parser_Command_def___closed__9; +x_1 = l_Lean_Parser_Command_definition___closed__5; +x_2 = l_Lean_Parser_Command_definition___closed__9; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_def___closed__11() { +static lean_object* _init_l_Lean_Parser_Command_definition___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_definition___closed__2; +x_2 = lean_unsigned_to_nat(1024u); +x_3 = l_Lean_Parser_Command_definition___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_definition___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Command_definition___closed__3; +x_2 = l_Lean_Parser_Command_definition___closed__11; +x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Command_definition___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Command_definition___closed__2; +x_2 = l_Lean_Parser_Command_definition___closed__12; +x_3 = l_Lean_Parser_withCache(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Command_definition() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Parser_Command_definition___closed__13; +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Command_def___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("def", 3); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Command_def___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_def___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_Command_def___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; +x_1 = l_Lean_Parser_Command_def___closed__1; +x_2 = l_Lean_Parser_Command_def___closed__2; +x_3 = 1; +x_4 = 0; +x_5 = l_Lean_Parser_mkAntiquot(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Parser_Command_def___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_def___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_def___closed__10; +x_3 = l_Lean_Parser_Command_definition___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_def___closed__12() { +static lean_object* _init_l_Lean_Parser_Command_def___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_def___closed__3; -x_2 = l_Lean_Parser_Command_def___closed__11; +x_2 = l_Lean_Parser_Command_def___closed__4; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_def___closed__13() { +static lean_object* _init_l_Lean_Parser_Command_def___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_def___closed__2; -x_2 = l_Lean_Parser_Command_def___closed__12; +x_2 = l_Lean_Parser_Command_def___closed__5; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -8389,7 +8482,7 @@ static lean_object* _init_l_Lean_Parser_Command_def() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Command_def___closed__13; +x_1 = l_Lean_Parser_Command_def___closed__6; return x_1; } } @@ -8456,7 +8549,7 @@ static lean_object* _init_l_Lean_Parser_Command_theorem___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_def___closed__6; +x_1 = l_Lean_Parser_Command_definition___closed__6; x_2 = l_Lean_Parser_Command_theorem___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -8583,7 +8676,7 @@ static lean_object* _init_l_Lean_Parser_Command_opaque___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_def___closed__6; +x_1 = l_Lean_Parser_Command_definition___closed__6; x_2 = l_Lean_Parser_Command_opaque___closed__7; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -8813,7 +8906,7 @@ static lean_object* _init_l_Lean_Parser_Command_axiom___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_def___closed__6; +x_1 = l_Lean_Parser_Command_definition___closed__6; x_2 = l_Lean_Parser_Command_declSig; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -9728,7 +9821,7 @@ static lean_object* _init_l_Lean_Parser_Command_inductive___closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_def___closed__6; +x_1 = l_Lean_Parser_Command_definition___closed__6; x_2 = l_Lean_Parser_Command_inductive___closed__16; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -9895,7 +9988,7 @@ static lean_object* _init_l_Lean_Parser_Command_classInductive___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_def___closed__6; +x_1 = l_Lean_Parser_Command_definition___closed__6; x_2 = l_Lean_Parser_Command_classInductive___closed__11; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -11437,7 +11530,7 @@ static lean_object* _init_l_Lean_Parser_Command_declaration___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_abbrev; +x_1 = l_Lean_Parser_Command_definition; x_2 = l_Lean_Parser_Command_declaration___closed__12; x_3 = l_Lean_Parser_orelse(x_1, x_2); return x_3; @@ -11447,39 +11540,49 @@ static lean_object* _init_l_Lean_Parser_Command_declaration___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_declaration___closed__4; +x_1 = l_Lean_Parser_Command_abbrev; x_2 = l_Lean_Parser_Command_declaration___closed__13; -x_3 = l_Lean_Parser_andthen(x_1, x_2); +x_3 = l_Lean_Parser_orelse(x_1, x_2); return x_3; } } static lean_object* _init_l_Lean_Parser_Command_declaration___closed__15() { _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_declaration___closed__14; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Command_declaration___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_declaration___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_declaration___closed__14; +x_3 = l_Lean_Parser_Command_declaration___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_declaration___closed__16() { +static lean_object* _init_l_Lean_Parser_Command_declaration___closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_declaration___closed__3; -x_2 = l_Lean_Parser_Command_declaration___closed__15; +x_2 = l_Lean_Parser_Command_declaration___closed__16; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_declaration___closed__17() { +static lean_object* _init_l_Lean_Parser_Command_declaration___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_declaration___closed__2; -x_2 = l_Lean_Parser_Command_declaration___closed__16; +x_2 = l_Lean_Parser_Command_declaration___closed__17; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -11488,7 +11591,7 @@ static lean_object* _init_l_Lean_Parser_Command_declaration() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Command_declaration___closed__17; +x_1 = l_Lean_Parser_Command_declaration___closed__18; return x_1; } } @@ -11509,7 +11612,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(199u); +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); @@ -11521,7 +11624,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(202u); +x_1 = lean_unsigned_to_nat(204u); x_2 = lean_unsigned_to_nat(50u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11549,7 +11652,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(199u); +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); @@ -11561,7 +11664,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(199u); +x_1 = lean_unsigned_to_nat(201u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13558,12 +13661,12 @@ 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_Command_def_formatter___closed__1() { +static lean_object* _init_l_Lean_Parser_Command_definition_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_def___closed__1; -x_2 = l_Lean_Parser_Command_def___closed__2; +x_1 = l_Lean_Parser_Command_definition___closed__1; +x_2 = l_Lean_Parser_Command_definition___closed__2; x_3 = 1; x_4 = 0; x_5 = lean_box(x_3); @@ -13576,17 +13679,17 @@ lean_closure_set(x_7, 3, x_6); return x_7; } } -static lean_object* _init_l_Lean_Parser_Command_def_formatter___closed__2() { +static lean_object* _init_l_Lean_Parser_Command_definition_formatter___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_def___closed__4; +x_1 = l_Lean_Parser_Command_definition___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_def_formatter___closed__3() { +static lean_object* _init_l_Lean_Parser_Command_definition_formatter___closed__3() { _start: { lean_object* x_1; lean_object* x_2; @@ -13596,7 +13699,7 @@ lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_def_formatter___closed__4() { +static lean_object* _init_l_Lean_Parser_Command_definition_formatter___closed__4() { _start: { lean_object* x_1; @@ -13604,61 +13707,136 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Command_optDefDeriving_formatter) return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_def_formatter___closed__5() { +static lean_object* _init_l_Lean_Parser_Command_definition_formatter___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_abbrev_formatter___closed__4; -x_2 = l_Lean_Parser_Command_def_formatter___closed__4; +x_2 = l_Lean_Parser_Command_definition_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_def_formatter___closed__6() { +static lean_object* _init_l_Lean_Parser_Command_definition_formatter___closed__6() { _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_def_formatter___closed__5; +x_2 = l_Lean_Parser_Command_definition_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_def_formatter___closed__7() { +static lean_object* _init_l_Lean_Parser_Command_definition_formatter___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_def_formatter___closed__3; -x_2 = l_Lean_Parser_Command_def_formatter___closed__6; +x_1 = l_Lean_Parser_Command_definition_formatter___closed__3; +x_2 = l_Lean_Parser_Command_definition_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_def_formatter___closed__8() { +static lean_object* _init_l_Lean_Parser_Command_definition_formatter___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_def_formatter___closed__2; -x_2 = l_Lean_Parser_Command_def_formatter___closed__7; +x_1 = l_Lean_Parser_Command_definition_formatter___closed__2; +x_2 = l_Lean_Parser_Command_definition_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_def_formatter___closed__9() { +static lean_object* _init_l_Lean_Parser_Command_definition_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_Command_definition___closed__2; +x_2 = lean_unsigned_to_nat(1024u); +x_3 = l_Lean_Parser_Command_definition_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); +lean_closure_set(x_4, 2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Parser_Command_definition_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_Command_definition_formatter___closed__1; +x_7 = l_Lean_Parser_Command_definition_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; +} +} +static lean_object* _init_l___regBuiltin_Lean_Parser_Command_definition_formatter___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_Parser_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_definition___closed__1; +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; +} +} +static lean_object* _init_l___regBuiltin_Lean_Parser_Command_definition_formatter___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Command_definition_formatter), 5, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_definition_formatter(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_2 = l___regBuiltin_Lean_Parser_Term_quot_formatter___closed__3; +x_3 = l_Lean_Parser_Command_definition___closed__2; +x_4 = l___regBuiltin_Lean_Parser_Command_definition_formatter___closed__1; +x_5 = l___regBuiltin_Lean_Parser_Command_definition_formatter___closed__2; +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_def_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_def___closed__1; +x_2 = l_Lean_Parser_Command_def___closed__2; +x_3 = 1; +x_4 = 0; +x_5 = lean_box(x_3); +x_6 = lean_box(x_4); +x_7 = lean_alloc_closure((void*)(l_Lean_Parser_mkAntiquot_formatter___boxed), 9, 4); +lean_closure_set(x_7, 0, x_1); +lean_closure_set(x_7, 1, x_2); +lean_closure_set(x_7, 2, x_5); +lean_closure_set(x_7, 3, x_6); +return x_7; +} +} +static lean_object* _init_l_Lean_Parser_Command_def_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_def___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_def_formatter___closed__8; +x_3 = l_Lean_Parser_Command_definition_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); @@ -13671,7 +13849,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_def_formatter(lean_object* x_1, l { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_def_formatter___closed__1; -x_7 = l_Lean_Parser_Command_def_formatter___closed__9; +x_7 = l_Lean_Parser_Command_def_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; } @@ -13858,7 +14036,7 @@ static lean_object* _init_l_Lean_Parser_Command_theorem_formatter___closed__5() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_def_formatter___closed__3; +x_1 = l_Lean_Parser_Command_definition_formatter___closed__3; x_2 = l_Lean_Parser_Command_theorem_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); @@ -13989,7 +14167,7 @@ static lean_object* _init_l_Lean_Parser_Command_opaque_formatter___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_def_formatter___closed__3; +x_1 = l_Lean_Parser_Command_definition_formatter___closed__3; x_2 = l_Lean_Parser_Command_opaque_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); @@ -14462,7 +14640,7 @@ static lean_object* _init_l_Lean_Parser_Command_axiom_formatter___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_def_formatter___closed__3; +x_1 = l_Lean_Parser_Command_definition_formatter___closed__3; x_2 = l_Lean_Parser_Command_theorem_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); @@ -15425,7 +15603,7 @@ static lean_object* _init_l_Lean_Parser_Command_inductive_formatter___closed__13 _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_def_formatter___closed__3; +x_1 = l_Lean_Parser_Command_definition_formatter___closed__3; x_2 = l_Lean_Parser_Command_inductive_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); @@ -15602,7 +15780,7 @@ static lean_object* _init_l_Lean_Parser_Command_classInductive_formatter___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_def_formatter___closed__3; +x_1 = l_Lean_Parser_Command_definition_formatter___closed__3; x_2 = l_Lean_Parser_Command_classInductive_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); @@ -17269,7 +17447,7 @@ static lean_object* _init_l_Lean_Parser_Command_declaration_formatter___closed__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Parser_Command_abbrev_formatter___closed__2; +x_1 = l___regBuiltin_Lean_Parser_Command_definition_formatter___closed__2; x_2 = l_Lean_Parser_Command_declaration_formatter___closed__10; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_orelse_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -17281,9 +17459,9 @@ static lean_object* _init_l_Lean_Parser_Command_declaration_formatter___closed__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_declaration_formatter___closed__2; +x_1 = l___regBuiltin_Lean_Parser_Command_abbrev_formatter___closed__2; x_2 = l_Lean_Parser_Command_declaration_formatter___closed__11; -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; @@ -17292,10 +17470,22 @@ return x_3; static lean_object* _init_l_Lean_Parser_Command_declaration_formatter___closed__13() { _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_declaration_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_declaration_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_declaration___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_declaration_formatter___closed__12; +x_3 = l_Lean_Parser_Command_declaration_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); @@ -17308,7 +17498,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_declaration_formatter(lean_object { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_declaration_formatter___closed__1; -x_7 = l_Lean_Parser_Command_declaration_formatter___closed__13; +x_7 = l_Lean_Parser_Command_declaration_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; } @@ -19300,12 +19490,12 @@ 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_Command_def_parenthesizer___closed__1() { +static lean_object* _init_l_Lean_Parser_Command_definition_parenthesizer___closed__1() { _start: { lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_1 = l_Lean_Parser_Command_def___closed__1; -x_2 = l_Lean_Parser_Command_def___closed__2; +x_1 = l_Lean_Parser_Command_definition___closed__1; +x_2 = l_Lean_Parser_Command_definition___closed__2; x_3 = 1; x_4 = 0; x_5 = lean_box(x_3); @@ -19318,17 +19508,17 @@ lean_closure_set(x_7, 3, x_6); return x_7; } } -static lean_object* _init_l_Lean_Parser_Command_def_parenthesizer___closed__2() { +static lean_object* _init_l_Lean_Parser_Command_definition_parenthesizer___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_def___closed__4; +x_1 = l_Lean_Parser_Command_definition___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_def_parenthesizer___closed__3() { +static lean_object* _init_l_Lean_Parser_Command_definition_parenthesizer___closed__3() { _start: { lean_object* x_1; lean_object* x_2; @@ -19338,7 +19528,7 @@ lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_def_parenthesizer___closed__4() { +static lean_object* _init_l_Lean_Parser_Command_definition_parenthesizer___closed__4() { _start: { lean_object* x_1; @@ -19346,61 +19536,136 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Command_optDefDeriving_parenthesi return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_def_parenthesizer___closed__5() { +static lean_object* _init_l_Lean_Parser_Command_definition_parenthesizer___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_abbrev_parenthesizer___closed__4; -x_2 = l_Lean_Parser_Command_def_parenthesizer___closed__4; +x_2 = l_Lean_Parser_Command_definition_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_def_parenthesizer___closed__6() { +static lean_object* _init_l_Lean_Parser_Command_definition_parenthesizer___closed__6() { _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_def_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Command_definition_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_def_parenthesizer___closed__7() { +static lean_object* _init_l_Lean_Parser_Command_definition_parenthesizer___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_def_parenthesizer___closed__3; -x_2 = l_Lean_Parser_Command_def_parenthesizer___closed__6; +x_1 = l_Lean_Parser_Command_definition_parenthesizer___closed__3; +x_2 = l_Lean_Parser_Command_definition_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_def_parenthesizer___closed__8() { +static lean_object* _init_l_Lean_Parser_Command_definition_parenthesizer___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_def_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Command_def_parenthesizer___closed__7; +x_1 = l_Lean_Parser_Command_definition_parenthesizer___closed__2; +x_2 = l_Lean_Parser_Command_definition_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_def_parenthesizer___closed__9() { +static lean_object* _init_l_Lean_Parser_Command_definition_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_Command_definition___closed__2; +x_2 = lean_unsigned_to_nat(1024u); +x_3 = l_Lean_Parser_Command_definition_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); +lean_closure_set(x_4, 2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Parser_Command_definition_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_Command_definition_parenthesizer___closed__1; +x_7 = l_Lean_Parser_Command_definition_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; +} +} +static lean_object* _init_l___regBuiltin_Lean_Parser_Command_definition_parenthesizer___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_Parser_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_definition___closed__1; +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; +} +} +static lean_object* _init_l___regBuiltin_Lean_Parser_Command_definition_parenthesizer___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Command_definition_parenthesizer), 5, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_definition_parenthesizer(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_2 = l___regBuiltin_Lean_Parser_Term_quot_parenthesizer___closed__3; +x_3 = l_Lean_Parser_Command_definition___closed__2; +x_4 = l___regBuiltin_Lean_Parser_Command_definition_parenthesizer___closed__1; +x_5 = l___regBuiltin_Lean_Parser_Command_definition_parenthesizer___closed__2; +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_def_parenthesizer___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_1 = l_Lean_Parser_Command_def___closed__1; +x_2 = l_Lean_Parser_Command_def___closed__2; +x_3 = 1; +x_4 = 0; +x_5 = lean_box(x_3); +x_6 = lean_box(x_4); +x_7 = lean_alloc_closure((void*)(l_Lean_Parser_mkAntiquot_parenthesizer___boxed), 9, 4); +lean_closure_set(x_7, 0, x_1); +lean_closure_set(x_7, 1, x_2); +lean_closure_set(x_7, 2, x_5); +lean_closure_set(x_7, 3, x_6); +return x_7; +} +} +static lean_object* _init_l_Lean_Parser_Command_def_parenthesizer___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_def___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_def_parenthesizer___closed__8; +x_3 = l_Lean_Parser_Command_definition_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); @@ -19413,7 +19678,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_def_parenthesizer(lean_object* x_ { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_def_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Command_def_parenthesizer___closed__9; +x_7 = l_Lean_Parser_Command_def_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; } @@ -19600,7 +19865,7 @@ static lean_object* _init_l_Lean_Parser_Command_theorem_parenthesizer___closed__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_def_parenthesizer___closed__3; +x_1 = l_Lean_Parser_Command_definition_parenthesizer___closed__3; x_2 = l_Lean_Parser_Command_theorem_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); @@ -19731,7 +19996,7 @@ static lean_object* _init_l_Lean_Parser_Command_opaque_parenthesizer___closed__5 _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_def_parenthesizer___closed__3; +x_1 = l_Lean_Parser_Command_definition_parenthesizer___closed__3; x_2 = l_Lean_Parser_Command_opaque_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); @@ -20206,7 +20471,7 @@ static lean_object* _init_l_Lean_Parser_Command_axiom_parenthesizer___closed__3( _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_def_parenthesizer___closed__3; +x_1 = l_Lean_Parser_Command_definition_parenthesizer___closed__3; x_2 = l_Lean_Parser_Command_theorem_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); @@ -21169,7 +21434,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_Lean_Parser_Command_def_parenthesizer___closed__3; +x_1 = l_Lean_Parser_Command_definition_parenthesizer___closed__3; x_2 = l_Lean_Parser_Command_inductive_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); @@ -21346,7 +21611,7 @@ static lean_object* _init_l_Lean_Parser_Command_classInductive_parenthesizer___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_def_parenthesizer___closed__3; +x_1 = l_Lean_Parser_Command_definition_parenthesizer___closed__3; x_2 = l_Lean_Parser_Command_classInductive_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); @@ -23013,7 +23278,7 @@ static lean_object* _init_l_Lean_Parser_Command_declaration_parenthesizer___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Parser_Command_abbrev_parenthesizer___closed__2; +x_1 = l___regBuiltin_Lean_Parser_Command_definition_parenthesizer___closed__2; x_2 = l_Lean_Parser_Command_declaration_parenthesizer___closed__10; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_orelse_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -23025,9 +23290,9 @@ static lean_object* _init_l_Lean_Parser_Command_declaration_parenthesizer___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_declaration_parenthesizer___closed__2; +x_1 = l___regBuiltin_Lean_Parser_Command_abbrev_parenthesizer___closed__2; x_2 = l_Lean_Parser_Command_declaration_parenthesizer___closed__11; -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; @@ -23036,10 +23301,22 @@ return x_3; static lean_object* _init_l_Lean_Parser_Command_declaration_parenthesizer___closed__13() { _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_declaration_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_declaration_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_declaration___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_declaration_parenthesizer___closed__12; +x_3 = l_Lean_Parser_Command_declaration_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); @@ -23052,7 +23329,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_declaration_parenthesizer(lean_ob { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_declaration_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Command_declaration_parenthesizer___closed__13; +x_7 = l_Lean_Parser_Command_declaration_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; } @@ -23274,7 +23551,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(203u); +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); @@ -23286,7 +23563,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(204u); +x_1 = lean_unsigned_to_nat(206u); x_2 = lean_unsigned_to_nat(94u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23314,7 +23591,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(203u); +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); @@ -23326,7 +23603,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(203u); +x_1 = lean_unsigned_to_nat(205u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23802,7 +24079,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(205u); +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); @@ -23814,7 +24091,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(206u); +x_1 = lean_unsigned_to_nat(208u); x_2 = lean_unsigned_to_nat(62u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23842,7 +24119,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(205u); +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); @@ -23854,7 +24131,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(205u); +x_1 = lean_unsigned_to_nat(207u); x_2 = lean_unsigned_to_nat(50u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -24242,7 +24519,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(207u); +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); @@ -24254,7 +24531,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(208u); +x_1 = lean_unsigned_to_nat(210u); x_2 = lean_unsigned_to_nat(42u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -24282,7 +24559,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(207u); +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); @@ -24294,7 +24571,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(207u); +x_1 = lean_unsigned_to_nat(209u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -24605,7 +24882,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(209u); +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); @@ -24617,7 +24894,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(210u); +x_1 = lean_unsigned_to_nat(212u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -24645,7 +24922,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(209u); +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); @@ -24657,7 +24934,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(209u); +x_1 = lean_unsigned_to_nat(211u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -25004,7 +25281,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(211u); +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); @@ -25016,7 +25293,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(212u); +x_1 = lean_unsigned_to_nat(214u); x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -25044,7 +25321,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(211u); +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); @@ -25056,7 +25333,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(211u); +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); @@ -25412,7 +25689,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(213u); +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); @@ -25424,7 +25701,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(214u); +x_1 = lean_unsigned_to_nat(216u); x_2 = lean_unsigned_to_nat(55u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -25452,7 +25729,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(213u); +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); @@ -25464,7 +25741,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(213u); +x_1 = lean_unsigned_to_nat(215u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -25840,7 +26117,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(215u); +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); @@ -25852,7 +26129,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(216u); +x_1 = lean_unsigned_to_nat(218u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -25880,7 +26157,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(215u); +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); @@ -25892,7 +26169,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(215u); +x_1 = lean_unsigned_to_nat(217u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -26267,7 +26544,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(217u); +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); @@ -26279,7 +26556,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(218u); +x_1 = lean_unsigned_to_nat(220u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -26307,7 +26584,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(217u); +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); @@ -26319,7 +26596,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(217u); +x_1 = lean_unsigned_to_nat(219u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -26674,7 +26951,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(219u); +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); @@ -26686,7 +26963,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(220u); +x_1 = lean_unsigned_to_nat(222u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -26714,7 +26991,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(219u); +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); @@ -26726,7 +27003,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(219u); +x_1 = lean_unsigned_to_nat(221u); x_2 = lean_unsigned_to_nat(43u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27081,7 +27358,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(221u); +x_1 = lean_unsigned_to_nat(223u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27093,7 +27370,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(222u); +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); @@ -27120,7 +27397,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(221u); +x_1 = lean_unsigned_to_nat(223u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27132,7 +27409,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(221u); +x_1 = lean_unsigned_to_nat(223u); x_2 = lean_unsigned_to_nat(36u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27487,7 +27764,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(223u); +x_1 = lean_unsigned_to_nat(225u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27499,7 +27776,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(224u); +x_1 = lean_unsigned_to_nat(226u); x_2 = lean_unsigned_to_nat(24u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27527,7 +27804,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(223u); +x_1 = lean_unsigned_to_nat(225u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27539,7 +27816,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(223u); +x_1 = lean_unsigned_to_nat(225u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27894,7 +28171,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(225u); +x_1 = lean_unsigned_to_nat(227u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27906,7 +28183,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(226u); +x_1 = lean_unsigned_to_nat(228u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27934,7 +28211,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(225u); +x_1 = lean_unsigned_to_nat(227u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27946,7 +28223,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(225u); +x_1 = lean_unsigned_to_nat(227u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28291,7 +28568,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(227u); +x_1 = lean_unsigned_to_nat(229u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28303,7 +28580,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(228u); +x_1 = lean_unsigned_to_nat(230u); x_2 = lean_unsigned_to_nat(9u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28331,7 +28608,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(227u); +x_1 = lean_unsigned_to_nat(229u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28343,7 +28620,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(227u); +x_1 = lean_unsigned_to_nat(229u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28684,7 +28961,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(229u); +x_1 = lean_unsigned_to_nat(231u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28696,7 +28973,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(230u); +x_1 = lean_unsigned_to_nat(232u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28724,7 +29001,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(229u); +x_1 = lean_unsigned_to_nat(231u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28736,7 +29013,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(229u); +x_1 = lean_unsigned_to_nat(231u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29142,7 +29419,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(231u); +x_1 = lean_unsigned_to_nat(233u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29154,7 +29431,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(232u); +x_1 = lean_unsigned_to_nat(234u); x_2 = lean_unsigned_to_nat(51u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29182,7 +29459,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(231u); +x_1 = lean_unsigned_to_nat(233u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29194,7 +29471,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(231u); +x_1 = lean_unsigned_to_nat(233u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29618,7 +29895,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_printEqns_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(235u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29630,7 +29907,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_printEqns_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(236u); x_2 = lean_unsigned_to_nat(86u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29658,7 +29935,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_printEqns_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(235u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29670,7 +29947,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_printEqns_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(235u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30087,7 +30364,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(235u); +x_1 = lean_unsigned_to_nat(237u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30099,7 +30376,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(236u); +x_1 = lean_unsigned_to_nat(238u); x_2 = lean_unsigned_to_nat(13u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30127,7 +30404,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(235u); +x_1 = lean_unsigned_to_nat(237u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30139,7 +30416,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(235u); +x_1 = lean_unsigned_to_nat(237u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30564,7 +30841,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(238u); +x_1 = lean_unsigned_to_nat(240u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30576,7 +30853,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(239u); +x_1 = lean_unsigned_to_nat(241u); x_2 = lean_unsigned_to_nat(50u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30604,7 +30881,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(238u); +x_1 = lean_unsigned_to_nat(240u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30616,7 +30893,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(238u); +x_1 = lean_unsigned_to_nat(240u); x_2 = lean_unsigned_to_nat(42u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -31331,7 +31608,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(242u); +x_1 = lean_unsigned_to_nat(244u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -31343,7 +31620,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(245u); +x_1 = lean_unsigned_to_nat(247u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -31371,7 +31648,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(242u); +x_1 = lean_unsigned_to_nat(244u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -31383,7 +31660,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(242u); +x_1 = lean_unsigned_to_nat(244u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -32128,7 +32405,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(246u); +x_1 = lean_unsigned_to_nat(248u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -32140,7 +32417,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(247u); +x_1 = lean_unsigned_to_nat(249u); x_2 = lean_unsigned_to_nat(50u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -32168,7 +32445,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(246u); +x_1 = lean_unsigned_to_nat(248u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -32180,7 +32457,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(246u); +x_1 = lean_unsigned_to_nat(248u); x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -32589,7 +32866,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(248u); +x_1 = lean_unsigned_to_nat(250u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -32601,7 +32878,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(249u); +x_1 = lean_unsigned_to_nat(251u); x_2 = lean_unsigned_to_nat(10u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -32629,7 +32906,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(248u); +x_1 = lean_unsigned_to_nat(250u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -32641,7 +32918,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(248u); +x_1 = lean_unsigned_to_nat(250u); x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -33770,7 +34047,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(266u); +x_1 = lean_unsigned_to_nat(268u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -33782,7 +34059,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(267u); +x_1 = lean_unsigned_to_nat(269u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -33810,7 +34087,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(266u); +x_1 = lean_unsigned_to_nat(268u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -33822,7 +34099,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(266u); +x_1 = lean_unsigned_to_nat(268u); x_2 = lean_unsigned_to_nat(36u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -35864,7 +36141,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(269u); +x_1 = lean_unsigned_to_nat(271u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -35876,7 +36153,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(271u); +x_1 = lean_unsigned_to_nat(273u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -35904,7 +36181,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(269u); +x_1 = lean_unsigned_to_nat(271u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -35916,7 +36193,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(269u); +x_1 = lean_unsigned_to_nat(271u); x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -36593,7 +36870,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(274u); +x_1 = lean_unsigned_to_nat(276u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -36605,7 +36882,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(276u); +x_1 = lean_unsigned_to_nat(278u); x_2 = lean_unsigned_to_nat(87u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -36633,7 +36910,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(274u); +x_1 = lean_unsigned_to_nat(276u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -36645,7 +36922,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(274u); +x_1 = lean_unsigned_to_nat(276u); x_2 = lean_unsigned_to_nat(42u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -37383,7 +37660,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(278u); +x_1 = lean_unsigned_to_nat(280u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -37395,7 +37672,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(279u); +x_1 = lean_unsigned_to_nat(281u); x_2 = lean_unsigned_to_nat(47u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -37423,7 +37700,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(278u); +x_1 = lean_unsigned_to_nat(280u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -37435,7 +37712,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(278u); +x_1 = lean_unsigned_to_nat(280u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -37780,7 +38057,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_addDocString_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(281u); +x_1 = lean_unsigned_to_nat(283u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -37792,7 +38069,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_addDocString_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(282u); +x_1 = lean_unsigned_to_nat(284u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -37820,7 +38097,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_addDocString_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(281u); +x_1 = lean_unsigned_to_nat(283u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -37832,7 +38109,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_addDocString_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(281u); +x_1 = lean_unsigned_to_nat(283u); x_2 = lean_unsigned_to_nat(42u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -38229,7 +38506,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_deriveInduction_dec _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(292u); +x_1 = lean_unsigned_to_nat(294u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -38241,7 +38518,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_deriveInduction_dec _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(293u); +x_1 = lean_unsigned_to_nat(295u); x_2 = lean_unsigned_to_nat(48u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -38269,7 +38546,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_deriveInduction_dec _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(292u); +x_1 = lean_unsigned_to_nat(294u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -38281,7 +38558,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_deriveInduction_dec _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(292u); +x_1 = lean_unsigned_to_nat(294u); x_2 = lean_unsigned_to_nat(45u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -38654,7 +38931,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_genInjectiveTheorem _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(299u); +x_1 = lean_unsigned_to_nat(301u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -38666,7 +38943,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_genInjectiveTheorem _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(300u); +x_1 = lean_unsigned_to_nat(302u); x_2 = lean_unsigned_to_nat(37u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -38694,7 +38971,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_genInjectiveTheorem _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(299u); +x_1 = lean_unsigned_to_nat(301u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -38706,7 +38983,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_genInjectiveTheorem _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(299u); +x_1 = lean_unsigned_to_nat(301u); x_2 = lean_unsigned_to_nat(50u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -39204,7 +39481,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_2791_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2829_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; @@ -39265,7 +39542,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_2813____closed__1() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -39275,7 +39552,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_2813____closed__2() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__2() { _start: { lean_object* x_1; @@ -39283,19 +39560,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_2813____closed__3() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____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_2813____closed__2; +x_4 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____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_2813____closed__4() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__4() { _start: { lean_object* x_1; lean_object* x_2; @@ -39305,7 +39582,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_2813____closed__5() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__5() { _start: { lean_object* x_1; lean_object* x_2; @@ -39315,7 +39592,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_2813____closed__6() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__6() { _start: { lean_object* x_1; lean_object* x_2; @@ -39325,12 +39602,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_2813____closed__7() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____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_2813____closed__6; +x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__6; x_3 = 1; x_4 = lean_alloc_ctor(0, 2, 1); lean_ctor_set(x_4, 0, x_1); @@ -39339,7 +39616,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_2813____closed__8() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__8() { _start: { lean_object* x_1; @@ -39347,17 +39624,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_2813____closed__9() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__9() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__8; +x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____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_2813____closed__10() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__10() { _start: { lean_object* x_1; @@ -39365,7 +39642,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_2813____closed__11() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__11() { _start: { lean_object* x_1; @@ -39373,17 +39650,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_2813____closed__12() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__12() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__11; +x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____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_2813____closed__13() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__13() { _start: { lean_object* x_1; @@ -39391,7 +39668,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_2813____closed__14() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__14() { _start: { lean_object* x_1; @@ -39399,17 +39676,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_2813____closed__15() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____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_2813____closed__14; +x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____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_2813____closed__16() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__16() { _start: { lean_object* x_1; @@ -39417,19 +39694,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_2813____closed__17() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____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_2813____closed__16; +x_4 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____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_2813____closed__18() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__18() { _start: { lean_object* x_1; lean_object* x_2; @@ -39439,7 +39716,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_2813____closed__19() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__19() { _start: { lean_object* x_1; @@ -39447,17 +39724,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_2813____closed__20() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__20() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__19; +x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____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_2813____closed__21() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__21() { _start: { lean_object* x_1; @@ -39465,17 +39742,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_2813____closed__22() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__22() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__21; +x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____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_2813____closed__23() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -39485,7 +39762,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_2813____closed__24() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__24() { _start: { lean_object* x_1; lean_object* x_2; @@ -39495,7 +39772,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_2813____closed__25() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__25() { _start: { lean_object* x_1; lean_object* x_2; @@ -39505,7 +39782,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_2813____closed__26() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__26() { _start: { lean_object* x_1; lean_object* x_2; @@ -39515,7 +39792,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_2813____closed__27() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__27() { _start: { lean_object* x_1; lean_object* x_2; @@ -39525,7 +39802,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_2813____closed__28() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__28() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -39535,7 +39812,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_2813____closed__29() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__29() { _start: { lean_object* x_1; lean_object* x_2; @@ -39545,7 +39822,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_2813____closed__30() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__30() { _start: { lean_object* x_1; lean_object* x_2; @@ -39555,7 +39832,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_2813____closed__31() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__31() { _start: { lean_object* x_1; lean_object* x_2; @@ -39565,7 +39842,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_2813____closed__32() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__32() { _start: { lean_object* x_1; lean_object* x_2; @@ -39575,7 +39852,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_2813____closed__33() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__33() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -39585,7 +39862,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_2813____closed__34() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__34() { _start: { lean_object* x_1; lean_object* x_2; @@ -39595,7 +39872,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_2813____closed__35() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__35() { _start: { lean_object* x_1; lean_object* x_2; @@ -39605,7 +39882,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_2813____closed__36() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__36() { _start: { lean_object* x_1; lean_object* x_2; @@ -39615,7 +39892,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_2813____closed__37() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__37() { _start: { lean_object* x_1; lean_object* x_2; @@ -39625,7 +39902,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_2813____closed__38() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__38() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -39635,7 +39912,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_2813____closed__39() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__39() { _start: { lean_object* x_1; lean_object* x_2; @@ -39645,7 +39922,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_2813____closed__40() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__40() { _start: { lean_object* x_1; lean_object* x_2; @@ -39655,7 +39932,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_2813____closed__41() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__41() { _start: { lean_object* x_1; lean_object* x_2; @@ -39665,7 +39942,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_2813____closed__42() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__42() { _start: { lean_object* x_1; lean_object* x_2; @@ -39675,7 +39952,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_2813____closed__43() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__43() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -39685,7 +39962,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_2813____closed__44() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__44() { _start: { lean_object* x_1; lean_object* x_2; @@ -39695,7 +39972,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_2813____closed__45() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__45() { _start: { lean_object* x_1; lean_object* x_2; @@ -39705,7 +39982,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_2813____closed__46() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__46() { _start: { lean_object* x_1; lean_object* x_2; @@ -39715,7 +39992,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_2813____closed__47() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__47() { _start: { lean_object* x_1; lean_object* x_2; @@ -39725,7 +40002,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_2813____closed__48() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__48() { _start: { lean_object* x_1; @@ -39733,29 +40010,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_2813____closed__49() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____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_2813____closed__48; +x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____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_2813____closed__50() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____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_2813____closed__48; +x_4 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____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_2813____closed__51() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__51() { _start: { lean_object* x_1; lean_object* x_2; @@ -39765,17 +40042,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_2813____closed__52() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__52() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__50; +x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____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_2813____closed__53() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__53() { _start: { lean_object* x_1; lean_object* x_2; @@ -39785,7 +40062,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_2813____closed__54() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__54() { _start: { lean_object* x_1; lean_object* x_2; @@ -39795,15 +40072,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_2813_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851_(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_2813____closed__1; -x_3 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__3; -x_4 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__4; -x_5 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__5; -x_6 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__7; +x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__1; +x_3 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__3; +x_4 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__4; +x_5 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__5; +x_6 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____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) { @@ -39811,8 +40088,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_2813____closed__10; -x_10 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__9; +x_9 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__10; +x_10 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__9; x_11 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_2, x_10, x_8); if (lean_obj_tag(x_11) == 0) { @@ -39820,8 +40097,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_2813____closed__13; -x_14 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__12; +x_13 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__13; +x_14 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__12; x_15 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_2, x_14, x_12); if (lean_obj_tag(x_15) == 0) { @@ -39829,9 +40106,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_2813____closed__15; -x_18 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__17; -x_19 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__18; +x_17 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__15; +x_18 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__17; +x_19 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____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) { @@ -39839,7 +40116,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_2813____closed__20; +x_22 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__20; x_23 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_17, x_22, x_21); if (lean_obj_tag(x_23) == 0) { @@ -39847,7 +40124,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_2813____closed__22; +x_25 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__22; x_26 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_17, x_25, x_24); if (lean_obj_tag(x_26) == 0) { @@ -39855,10 +40132,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_2813____closed__23; +x_28 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__23; x_29 = l_Lean_Parser_Command_declId___closed__2; -x_30 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__24; -x_31 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__25; +x_30 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__24; +x_31 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____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) { @@ -39866,7 +40143,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_2813____closed__26; +x_34 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__26; x_35 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_28, x_34, x_33); if (lean_obj_tag(x_35) == 0) { @@ -39874,7 +40151,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_2813____closed__27; +x_37 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__27; x_38 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_28, x_37, x_36); if (lean_obj_tag(x_38) == 0) { @@ -39882,10 +40159,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_2813____closed__28; +x_40 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__28; x_41 = l_Lean_Parser_Command_declSig___closed__2; -x_42 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__29; -x_43 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__30; +x_42 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__29; +x_43 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____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) { @@ -39893,7 +40170,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_2813____closed__31; +x_46 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__31; x_47 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_40, x_46, x_45); if (lean_obj_tag(x_47) == 0) { @@ -39901,7 +40178,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_2813____closed__32; +x_49 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__32; x_50 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_40, x_49, x_48); if (lean_obj_tag(x_50) == 0) { @@ -39909,10 +40186,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_2813____closed__33; +x_52 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__33; x_53 = l_Lean_Parser_Command_declVal___closed__2; -x_54 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__34; -x_55 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__35; +x_54 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__34; +x_55 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____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) { @@ -39920,7 +40197,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_2813____closed__36; +x_58 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__36; x_59 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_52, x_58, x_57); if (lean_obj_tag(x_59) == 0) { @@ -39928,7 +40205,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_2813____closed__37; +x_61 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__37; x_62 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_52, x_61, x_60); if (lean_obj_tag(x_62) == 0) { @@ -39936,10 +40213,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_2813____closed__38; +x_64 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__38; x_65 = l_Lean_Parser_Command_optDeclSig___closed__2; -x_66 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__39; -x_67 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__40; +x_66 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__39; +x_67 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____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) { @@ -39947,7 +40224,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_2813____closed__41; +x_70 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__41; x_71 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_64, x_70, x_69); if (lean_obj_tag(x_71) == 0) { @@ -39955,7 +40232,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_2813____closed__42; +x_73 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__42; x_74 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_64, x_73, x_72); if (lean_obj_tag(x_74) == 0) { @@ -39963,10 +40240,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_2813____closed__43; +x_76 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__43; x_77 = l_Lean_Parser_Command_openDecl___closed__2; -x_78 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__44; -x_79 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__45; +x_78 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__44; +x_79 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____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) { @@ -39974,7 +40251,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_2813____closed__46; +x_82 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__46; x_83 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_76, x_82, x_81); if (lean_obj_tag(x_83) == 0) { @@ -39982,7 +40259,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_2813____closed__47; +x_85 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__47; x_86 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_76, x_85, x_84); if (lean_obj_tag(x_86) == 0) { @@ -39990,10 +40267,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_2813____closed__49; -x_89 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__50; -x_90 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__51; -x_91 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__52; +x_88 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__49; +x_89 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__50; +x_90 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__51; +x_91 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____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) { @@ -40001,7 +40278,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_2813____closed__53; +x_94 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__53; x_95 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_88, x_94, x_93); if (lean_obj_tag(x_95) == 0) { @@ -40009,7 +40286,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_2813____closed__54; +x_97 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__54; x_98 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_88, x_97, x_96); return x_98; } @@ -40672,7 +40949,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_open_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(330u); +x_1 = lean_unsigned_to_nat(332u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -40684,7 +40961,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_open_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(331u); +x_1 = lean_unsigned_to_nat(333u); x_2 = lean_unsigned_to_nat(67u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -40712,7 +40989,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_open_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(330u); +x_1 = lean_unsigned_to_nat(332u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -40724,7 +41001,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_open_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(330u); +x_1 = lean_unsigned_to_nat(332u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -41150,7 +41427,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_set__option_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(337u); +x_1 = lean_unsigned_to_nat(339u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -41162,7 +41439,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_set__option_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(338u); +x_1 = lean_unsigned_to_nat(340u); x_2 = lean_unsigned_to_nat(82u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -41190,7 +41467,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_set__option_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(337u); +x_1 = lean_unsigned_to_nat(339u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -41202,7 +41479,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_set__option_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(337u); +x_1 = lean_unsigned_to_nat(339u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -41682,7 +41959,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_open_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(344u); +x_1 = lean_unsigned_to_nat(346u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -41694,7 +41971,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_open_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(345u); +x_1 = lean_unsigned_to_nat(347u); x_2 = lean_unsigned_to_nat(67u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -41722,7 +41999,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_open_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(344u); +x_1 = lean_unsigned_to_nat(346u); x_2 = lean_unsigned_to_nat(29u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -41734,7 +42011,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_open_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(344u); +x_1 = lean_unsigned_to_nat(346u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -42196,7 +42473,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_set__option_declRang _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(349u); +x_1 = lean_unsigned_to_nat(351u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -42208,7 +42485,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_set__option_declRang _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(350u); +x_1 = lean_unsigned_to_nat(352u); x_2 = lean_unsigned_to_nat(81u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -42236,7 +42513,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_set__option_declRang _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(349u); +x_1 = lean_unsigned_to_nat(351u); x_2 = lean_unsigned_to_nat(29u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -42248,7 +42525,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_set__option_declRang _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(349u); +x_1 = lean_unsigned_to_nat(351u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -43425,6 +43702,34 @@ l_Lean_Parser_Command_optDefDeriving___closed__11 = _init_l_Lean_Parser_Command_ lean_mark_persistent(l_Lean_Parser_Command_optDefDeriving___closed__11); l_Lean_Parser_Command_optDefDeriving = _init_l_Lean_Parser_Command_optDefDeriving(); lean_mark_persistent(l_Lean_Parser_Command_optDefDeriving); +l_Lean_Parser_Command_definition___closed__1 = _init_l_Lean_Parser_Command_definition___closed__1(); +lean_mark_persistent(l_Lean_Parser_Command_definition___closed__1); +l_Lean_Parser_Command_definition___closed__2 = _init_l_Lean_Parser_Command_definition___closed__2(); +lean_mark_persistent(l_Lean_Parser_Command_definition___closed__2); +l_Lean_Parser_Command_definition___closed__3 = _init_l_Lean_Parser_Command_definition___closed__3(); +lean_mark_persistent(l_Lean_Parser_Command_definition___closed__3); +l_Lean_Parser_Command_definition___closed__4 = _init_l_Lean_Parser_Command_definition___closed__4(); +lean_mark_persistent(l_Lean_Parser_Command_definition___closed__4); +l_Lean_Parser_Command_definition___closed__5 = _init_l_Lean_Parser_Command_definition___closed__5(); +lean_mark_persistent(l_Lean_Parser_Command_definition___closed__5); +l_Lean_Parser_Command_definition___closed__6 = _init_l_Lean_Parser_Command_definition___closed__6(); +lean_mark_persistent(l_Lean_Parser_Command_definition___closed__6); +l_Lean_Parser_Command_definition___closed__7 = _init_l_Lean_Parser_Command_definition___closed__7(); +lean_mark_persistent(l_Lean_Parser_Command_definition___closed__7); +l_Lean_Parser_Command_definition___closed__8 = _init_l_Lean_Parser_Command_definition___closed__8(); +lean_mark_persistent(l_Lean_Parser_Command_definition___closed__8); +l_Lean_Parser_Command_definition___closed__9 = _init_l_Lean_Parser_Command_definition___closed__9(); +lean_mark_persistent(l_Lean_Parser_Command_definition___closed__9); +l_Lean_Parser_Command_definition___closed__10 = _init_l_Lean_Parser_Command_definition___closed__10(); +lean_mark_persistent(l_Lean_Parser_Command_definition___closed__10); +l_Lean_Parser_Command_definition___closed__11 = _init_l_Lean_Parser_Command_definition___closed__11(); +lean_mark_persistent(l_Lean_Parser_Command_definition___closed__11); +l_Lean_Parser_Command_definition___closed__12 = _init_l_Lean_Parser_Command_definition___closed__12(); +lean_mark_persistent(l_Lean_Parser_Command_definition___closed__12); +l_Lean_Parser_Command_definition___closed__13 = _init_l_Lean_Parser_Command_definition___closed__13(); +lean_mark_persistent(l_Lean_Parser_Command_definition___closed__13); +l_Lean_Parser_Command_definition = _init_l_Lean_Parser_Command_definition(); +lean_mark_persistent(l_Lean_Parser_Command_definition); l_Lean_Parser_Command_def___closed__1 = _init_l_Lean_Parser_Command_def___closed__1(); lean_mark_persistent(l_Lean_Parser_Command_def___closed__1); l_Lean_Parser_Command_def___closed__2 = _init_l_Lean_Parser_Command_def___closed__2(); @@ -43437,20 +43742,6 @@ l_Lean_Parser_Command_def___closed__5 = _init_l_Lean_Parser_Command_def___closed lean_mark_persistent(l_Lean_Parser_Command_def___closed__5); l_Lean_Parser_Command_def___closed__6 = _init_l_Lean_Parser_Command_def___closed__6(); lean_mark_persistent(l_Lean_Parser_Command_def___closed__6); -l_Lean_Parser_Command_def___closed__7 = _init_l_Lean_Parser_Command_def___closed__7(); -lean_mark_persistent(l_Lean_Parser_Command_def___closed__7); -l_Lean_Parser_Command_def___closed__8 = _init_l_Lean_Parser_Command_def___closed__8(); -lean_mark_persistent(l_Lean_Parser_Command_def___closed__8); -l_Lean_Parser_Command_def___closed__9 = _init_l_Lean_Parser_Command_def___closed__9(); -lean_mark_persistent(l_Lean_Parser_Command_def___closed__9); -l_Lean_Parser_Command_def___closed__10 = _init_l_Lean_Parser_Command_def___closed__10(); -lean_mark_persistent(l_Lean_Parser_Command_def___closed__10); -l_Lean_Parser_Command_def___closed__11 = _init_l_Lean_Parser_Command_def___closed__11(); -lean_mark_persistent(l_Lean_Parser_Command_def___closed__11); -l_Lean_Parser_Command_def___closed__12 = _init_l_Lean_Parser_Command_def___closed__12(); -lean_mark_persistent(l_Lean_Parser_Command_def___closed__12); -l_Lean_Parser_Command_def___closed__13 = _init_l_Lean_Parser_Command_def___closed__13(); -lean_mark_persistent(l_Lean_Parser_Command_def___closed__13); l_Lean_Parser_Command_def = _init_l_Lean_Parser_Command_def(); lean_mark_persistent(l_Lean_Parser_Command_def); l_Lean_Parser_Command_theorem___closed__1 = _init_l_Lean_Parser_Command_theorem___closed__1(); @@ -44083,6 +44374,8 @@ l_Lean_Parser_Command_declaration___closed__16 = _init_l_Lean_Parser_Command_dec lean_mark_persistent(l_Lean_Parser_Command_declaration___closed__16); l_Lean_Parser_Command_declaration___closed__17 = _init_l_Lean_Parser_Command_declaration___closed__17(); lean_mark_persistent(l_Lean_Parser_Command_declaration___closed__17); +l_Lean_Parser_Command_declaration___closed__18 = _init_l_Lean_Parser_Command_declaration___closed__18(); +lean_mark_persistent(l_Lean_Parser_Command_declaration___closed__18); l_Lean_Parser_Command_declaration = _init_l_Lean_Parser_Command_declaration(); lean_mark_persistent(l_Lean_Parser_Command_declaration); if (builtin) {res = l___regBuiltin_Lean_Parser_Command_declaration(lean_io_mk_world()); @@ -44416,24 +44709,35 @@ l_Lean_Parser_Command_optDefDeriving_formatter___closed__6 = _init_l_Lean_Parser lean_mark_persistent(l_Lean_Parser_Command_optDefDeriving_formatter___closed__6); l_Lean_Parser_Command_optDefDeriving_formatter___closed__7 = _init_l_Lean_Parser_Command_optDefDeriving_formatter___closed__7(); lean_mark_persistent(l_Lean_Parser_Command_optDefDeriving_formatter___closed__7); -l_Lean_Parser_Command_def_formatter___closed__1 = _init_l_Lean_Parser_Command_def_formatter___closed__1(); +l_Lean_Parser_Command_definition_formatter___closed__1 = _init_l_Lean_Parser_Command_definition_formatter___closed__1(); +lean_mark_persistent(l_Lean_Parser_Command_definition_formatter___closed__1); +l_Lean_Parser_Command_definition_formatter___closed__2 = _init_l_Lean_Parser_Command_definition_formatter___closed__2(); +lean_mark_persistent(l_Lean_Parser_Command_definition_formatter___closed__2); +l_Lean_Parser_Command_definition_formatter___closed__3 = _init_l_Lean_Parser_Command_definition_formatter___closed__3(); +lean_mark_persistent(l_Lean_Parser_Command_definition_formatter___closed__3); +l_Lean_Parser_Command_definition_formatter___closed__4 = _init_l_Lean_Parser_Command_definition_formatter___closed__4(); +lean_mark_persistent(l_Lean_Parser_Command_definition_formatter___closed__4); +l_Lean_Parser_Command_definition_formatter___closed__5 = _init_l_Lean_Parser_Command_definition_formatter___closed__5(); +lean_mark_persistent(l_Lean_Parser_Command_definition_formatter___closed__5); +l_Lean_Parser_Command_definition_formatter___closed__6 = _init_l_Lean_Parser_Command_definition_formatter___closed__6(); +lean_mark_persistent(l_Lean_Parser_Command_definition_formatter___closed__6); +l_Lean_Parser_Command_definition_formatter___closed__7 = _init_l_Lean_Parser_Command_definition_formatter___closed__7(); +lean_mark_persistent(l_Lean_Parser_Command_definition_formatter___closed__7); +l_Lean_Parser_Command_definition_formatter___closed__8 = _init_l_Lean_Parser_Command_definition_formatter___closed__8(); +lean_mark_persistent(l_Lean_Parser_Command_definition_formatter___closed__8); +l_Lean_Parser_Command_definition_formatter___closed__9 = _init_l_Lean_Parser_Command_definition_formatter___closed__9(); +lean_mark_persistent(l_Lean_Parser_Command_definition_formatter___closed__9); +l___regBuiltin_Lean_Parser_Command_definition_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_definition_formatter___closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_definition_formatter___closed__1); +l___regBuiltin_Lean_Parser_Command_definition_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_definition_formatter___closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_definition_formatter___closed__2); +if (builtin) {res = l___regBuiltin_Lean_Parser_Command_definition_formatter(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}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(); lean_mark_persistent(l_Lean_Parser_Command_def_formatter___closed__2); -l_Lean_Parser_Command_def_formatter___closed__3 = _init_l_Lean_Parser_Command_def_formatter___closed__3(); -lean_mark_persistent(l_Lean_Parser_Command_def_formatter___closed__3); -l_Lean_Parser_Command_def_formatter___closed__4 = _init_l_Lean_Parser_Command_def_formatter___closed__4(); -lean_mark_persistent(l_Lean_Parser_Command_def_formatter___closed__4); -l_Lean_Parser_Command_def_formatter___closed__5 = _init_l_Lean_Parser_Command_def_formatter___closed__5(); -lean_mark_persistent(l_Lean_Parser_Command_def_formatter___closed__5); -l_Lean_Parser_Command_def_formatter___closed__6 = _init_l_Lean_Parser_Command_def_formatter___closed__6(); -lean_mark_persistent(l_Lean_Parser_Command_def_formatter___closed__6); -l_Lean_Parser_Command_def_formatter___closed__7 = _init_l_Lean_Parser_Command_def_formatter___closed__7(); -lean_mark_persistent(l_Lean_Parser_Command_def_formatter___closed__7); -l_Lean_Parser_Command_def_formatter___closed__8 = _init_l_Lean_Parser_Command_def_formatter___closed__8(); -lean_mark_persistent(l_Lean_Parser_Command_def_formatter___closed__8); -l_Lean_Parser_Command_def_formatter___closed__9 = _init_l_Lean_Parser_Command_def_formatter___closed__9(); -lean_mark_persistent(l_Lean_Parser_Command_def_formatter___closed__9); l___regBuiltin_Lean_Parser_Command_def_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_def_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_def_formatter___closed__1); l___regBuiltin_Lean_Parser_Command_def_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_def_formatter___closed__2(); @@ -45026,6 +45330,8 @@ l_Lean_Parser_Command_declaration_formatter___closed__12 = _init_l_Lean_Parser_C lean_mark_persistent(l_Lean_Parser_Command_declaration_formatter___closed__12); l_Lean_Parser_Command_declaration_formatter___closed__13 = _init_l_Lean_Parser_Command_declaration_formatter___closed__13(); lean_mark_persistent(l_Lean_Parser_Command_declaration_formatter___closed__13); +l_Lean_Parser_Command_declaration_formatter___closed__14 = _init_l_Lean_Parser_Command_declaration_formatter___closed__14(); +lean_mark_persistent(l_Lean_Parser_Command_declaration_formatter___closed__14); l___regBuiltin_Lean_Parser_Command_declaration_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_declaration_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_declaration_formatter___closed__1); l___regBuiltin_Lean_Parser_Command_declaration_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_declaration_formatter___closed__2(); @@ -45344,24 +45650,35 @@ l_Lean_Parser_Command_optDefDeriving_parenthesizer___closed__6 = _init_l_Lean_Pa lean_mark_persistent(l_Lean_Parser_Command_optDefDeriving_parenthesizer___closed__6); l_Lean_Parser_Command_optDefDeriving_parenthesizer___closed__7 = _init_l_Lean_Parser_Command_optDefDeriving_parenthesizer___closed__7(); lean_mark_persistent(l_Lean_Parser_Command_optDefDeriving_parenthesizer___closed__7); -l_Lean_Parser_Command_def_parenthesizer___closed__1 = _init_l_Lean_Parser_Command_def_parenthesizer___closed__1(); +l_Lean_Parser_Command_definition_parenthesizer___closed__1 = _init_l_Lean_Parser_Command_definition_parenthesizer___closed__1(); +lean_mark_persistent(l_Lean_Parser_Command_definition_parenthesizer___closed__1); +l_Lean_Parser_Command_definition_parenthesizer___closed__2 = _init_l_Lean_Parser_Command_definition_parenthesizer___closed__2(); +lean_mark_persistent(l_Lean_Parser_Command_definition_parenthesizer___closed__2); +l_Lean_Parser_Command_definition_parenthesizer___closed__3 = _init_l_Lean_Parser_Command_definition_parenthesizer___closed__3(); +lean_mark_persistent(l_Lean_Parser_Command_definition_parenthesizer___closed__3); +l_Lean_Parser_Command_definition_parenthesizer___closed__4 = _init_l_Lean_Parser_Command_definition_parenthesizer___closed__4(); +lean_mark_persistent(l_Lean_Parser_Command_definition_parenthesizer___closed__4); +l_Lean_Parser_Command_definition_parenthesizer___closed__5 = _init_l_Lean_Parser_Command_definition_parenthesizer___closed__5(); +lean_mark_persistent(l_Lean_Parser_Command_definition_parenthesizer___closed__5); +l_Lean_Parser_Command_definition_parenthesizer___closed__6 = _init_l_Lean_Parser_Command_definition_parenthesizer___closed__6(); +lean_mark_persistent(l_Lean_Parser_Command_definition_parenthesizer___closed__6); +l_Lean_Parser_Command_definition_parenthesizer___closed__7 = _init_l_Lean_Parser_Command_definition_parenthesizer___closed__7(); +lean_mark_persistent(l_Lean_Parser_Command_definition_parenthesizer___closed__7); +l_Lean_Parser_Command_definition_parenthesizer___closed__8 = _init_l_Lean_Parser_Command_definition_parenthesizer___closed__8(); +lean_mark_persistent(l_Lean_Parser_Command_definition_parenthesizer___closed__8); +l_Lean_Parser_Command_definition_parenthesizer___closed__9 = _init_l_Lean_Parser_Command_definition_parenthesizer___closed__9(); +lean_mark_persistent(l_Lean_Parser_Command_definition_parenthesizer___closed__9); +l___regBuiltin_Lean_Parser_Command_definition_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_definition_parenthesizer___closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_definition_parenthesizer___closed__1); +l___regBuiltin_Lean_Parser_Command_definition_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_definition_parenthesizer___closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_definition_parenthesizer___closed__2); +if (builtin) {res = l___regBuiltin_Lean_Parser_Command_definition_parenthesizer(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}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(); lean_mark_persistent(l_Lean_Parser_Command_def_parenthesizer___closed__2); -l_Lean_Parser_Command_def_parenthesizer___closed__3 = _init_l_Lean_Parser_Command_def_parenthesizer___closed__3(); -lean_mark_persistent(l_Lean_Parser_Command_def_parenthesizer___closed__3); -l_Lean_Parser_Command_def_parenthesizer___closed__4 = _init_l_Lean_Parser_Command_def_parenthesizer___closed__4(); -lean_mark_persistent(l_Lean_Parser_Command_def_parenthesizer___closed__4); -l_Lean_Parser_Command_def_parenthesizer___closed__5 = _init_l_Lean_Parser_Command_def_parenthesizer___closed__5(); -lean_mark_persistent(l_Lean_Parser_Command_def_parenthesizer___closed__5); -l_Lean_Parser_Command_def_parenthesizer___closed__6 = _init_l_Lean_Parser_Command_def_parenthesizer___closed__6(); -lean_mark_persistent(l_Lean_Parser_Command_def_parenthesizer___closed__6); -l_Lean_Parser_Command_def_parenthesizer___closed__7 = _init_l_Lean_Parser_Command_def_parenthesizer___closed__7(); -lean_mark_persistent(l_Lean_Parser_Command_def_parenthesizer___closed__7); -l_Lean_Parser_Command_def_parenthesizer___closed__8 = _init_l_Lean_Parser_Command_def_parenthesizer___closed__8(); -lean_mark_persistent(l_Lean_Parser_Command_def_parenthesizer___closed__8); -l_Lean_Parser_Command_def_parenthesizer___closed__9 = _init_l_Lean_Parser_Command_def_parenthesizer___closed__9(); -lean_mark_persistent(l_Lean_Parser_Command_def_parenthesizer___closed__9); l___regBuiltin_Lean_Parser_Command_def_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_def_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_def_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Command_def_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_def_parenthesizer___closed__2(); @@ -45954,6 +46271,8 @@ l_Lean_Parser_Command_declaration_parenthesizer___closed__12 = _init_l_Lean_Pars lean_mark_persistent(l_Lean_Parser_Command_declaration_parenthesizer___closed__12); l_Lean_Parser_Command_declaration_parenthesizer___closed__13 = _init_l_Lean_Parser_Command_declaration_parenthesizer___closed__13(); lean_mark_persistent(l_Lean_Parser_Command_declaration_parenthesizer___closed__13); +l_Lean_Parser_Command_declaration_parenthesizer___closed__14 = _init_l_Lean_Parser_Command_declaration_parenthesizer___closed__14(); +lean_mark_persistent(l_Lean_Parser_Command_declaration_parenthesizer___closed__14); l___regBuiltin_Lean_Parser_Command_declaration_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_declaration_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_declaration_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Command_declaration_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_declaration_parenthesizer___closed__2(); @@ -48743,122 +49062,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); -if (builtin) {res = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2791_(lean_io_mk_world()); +if (builtin) {res = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2829_(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_2813____closed__1 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__1(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__1); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__2 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__2(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__2); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__3 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__3(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__3); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__4 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__4(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__4); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__5 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__5(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__5); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__6 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__6(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__6); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__7 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__7(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__7); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__8 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__8(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__8); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__9 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__9(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__9); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__10 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__10(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__10); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__11 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__11(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__11); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__12 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__12(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__12); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__13 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__13(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__13); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__14 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__14(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__14); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__15 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__15(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__15); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__16 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__16(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__16); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__17 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__17(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__17); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__18 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__18(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__18); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__19 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__19(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__19); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__20 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__20(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__20); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__21 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__21(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__21); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__22 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__22(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__22); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__23 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__23(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__23); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__24 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__24(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__24); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__25 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__25(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__25); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__26 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__26(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__26); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__27 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__27(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__27); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__28 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__28(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__28); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__29 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__29(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__29); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__30 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__30(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__30); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__31 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__31(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__31); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__32 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__32(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__32); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__33 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__33(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__33); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__34 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__34(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__34); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__35 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__35(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__35); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__36 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__36(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__36); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__37 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__37(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__37); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__38 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__38(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__38); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__39 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__39(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__39); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__40 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__40(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__40); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__41 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__41(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__41); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__42 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__42(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__42); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__43 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__43(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__43); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__44 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__44(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__44); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__45 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__45(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__45); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__46 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__46(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__46); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__47 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__47(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__47); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__48 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__48(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__48); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__49 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__49(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__49); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__50 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__50(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__50); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__51 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__51(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__51); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__52 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__52(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__52); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__53 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__53(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__53); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__54 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__54(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__54); -if (builtin) {res = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813_(lean_io_mk_world()); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__1 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__1(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__1); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__2 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__2(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__2); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__3 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__3(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__3); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__4 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__4(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__4); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__5 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__5(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__5); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__6 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__6(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__6); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__7 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__7(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__7); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__8 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__8(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__8); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__9 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__9(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__9); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__10 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__10(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__10); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__11 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__11(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__11); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__12 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__12(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__12); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__13 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__13(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__13); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__14 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__14(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__14); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__15 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__15(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__15); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__16 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__16(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__16); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__17 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__17(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__17); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__18 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__18(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__18); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__19 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__19(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__19); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__20 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__20(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__20); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__21 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__21(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__21); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__22 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__22(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__22); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__23 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__23(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__23); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__24 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__24(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__24); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__25 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__25(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__25); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__26 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__26(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__26); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__27 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__27(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__27); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__28 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__28(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__28); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__29 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__29(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__29); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__30 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__30(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__30); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__31 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__31(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__31); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__32 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__32(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__32); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__33 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__33(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__33); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__34 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__34(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__34); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__35 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__35(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__35); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__36 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__36(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__36); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__37 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__37(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__37); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__38 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__38(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__38); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__39 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__39(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__39); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__40 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__40(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__40); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__41 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__41(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__41); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__42 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__42(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__42); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__43 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__43(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__43); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__44 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__44(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__44); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__45 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__45(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__45); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__46 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__46(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__46); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__47 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__47(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__47); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__48 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__48(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__48); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__49 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__49(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__49); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__50 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__50(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__50); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__51 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__51(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__51); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__52 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__52(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__52); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__53 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__53(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__53); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__54 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__54(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__54); +if (builtin) {res = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851_(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(); diff --git a/stage0/stdlib/Lean/PrettyPrinter/Delaborator/TopDownAnalyze.c b/stage0/stdlib/Lean/PrettyPrinter/Delaborator/TopDownAnalyze.c index 8befaec1e54c..fa852ab6b2cd 100644 --- a/stage0/stdlib/Lean/PrettyPrinter/Delaborator/TopDownAnalyze.c +++ b/stage0/stdlib/Lean/PrettyPrinter/Delaborator/TopDownAnalyze.c @@ -4747,7 +4747,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_replac lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_replaceLPsWithVars___lambda__1___closed__1; x_2 = l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_replaceLPsWithVars___lambda__1___closed__2; -x_3 = lean_unsigned_to_nat(184u); +x_3 = lean_unsigned_to_nat(210u); x_4 = lean_unsigned_to_nat(14u); x_5 = l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_replaceLPsWithVars___lambda__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/Replay.c b/stage0/stdlib/Lean/Replay.c index 53e6275a69d1..6e705baaee9c 100644 --- a/stage0/stdlib/Lean/Replay.c +++ b/stage0/stdlib/Lean/Replay.c @@ -1190,7 +1190,7 @@ static lean_object* _init_l_List_mapM_loop___at_Lean_Environment_Replay_replayCo 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_mapM_loop___at_Lean_Environment_Replay_replayConstant___spec__3___closed__1; x_2 = l_List_mapM_loop___at_Lean_Environment_Replay_replayConstant___spec__3___closed__2; -x_3 = lean_unsigned_to_nat(184u); +x_3 = lean_unsigned_to_nat(210u); x_4 = lean_unsigned_to_nat(14u); x_5 = l_List_mapM_loop___at_Lean_Environment_Replay_replayConstant___spec__3___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/Server/FileWorker/RequestHandling.c b/stage0/stdlib/Lean/Server/FileWorker/RequestHandling.c index afcdd1c48747..bd339e198393 100644 --- a/stage0/stdlib/Lean/Server/FileWorker/RequestHandling.c +++ b/stage0/stdlib/Lean/Server/FileWorker/RequestHandling.c @@ -19,27 +19,27 @@ static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymb lean_object* l_Lean_Name_getNumParts(lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at_Lean_Server_FileWorker_locationLinksOfInfo___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__8; +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__17___lambda__4___closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_NamespaceEntry_finish(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__27___lambda__1(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__21___lambda__2(lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__24___lambda__3___closed__1; lean_object* lean_format_pretty(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_InfoTree_goalsAt_x3f(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleHover___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Server_FileWorker_locationLinksOfInfo___spec__4___boxed(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_11383____spec__26___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__21; lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8890_(lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__33(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleHover___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__18(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__10___lambda__1(lean_object*); lean_object* l_Lean_Json_compress(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__30___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleHover(lean_object*, lean_object*, lean_object*); uint32_t lean_string_utf8_get(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__25(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_FileWorker_handleDocumentHighlight___spec__3___lambda__1(lean_object*, size_t, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__13; lean_object* l_Lean_Server_Snapshots_Snapshot_infoTree(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__37___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__7___lambda__1(lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__4___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleHover___lambda__6___closed__1; lean_object* l_Lean_Widget_diffInteractiveGoals(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__5; @@ -51,55 +51,48 @@ LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDefinition(uint8_t, lean lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); lean_object* l_Lean_Server_documentUriFromModule(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handlePlainGoal___lambda__1___closed__1; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__4___boxed(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_11385____spec__26___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleSemanticTokens_go___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addCommandRange___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleHover___lambda__2___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_handleDocumentHighlight___spec__1___boxed(lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9208_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Json_mkObj(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__1(lean_object*); static lean_object* l_Lean_Server_FileWorker_handleCompletion___closed__2; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_handlePlainGoal___spec__2___closed__1; LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___spec__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_List_mapM_loop___at_Lean_Server_FileWorker_getInteractiveGoals___spec__2___closed__2; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__24___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__7___lambda__3___closed__1; -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__30; -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__30___lambda__3___closed__1; +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__26; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__30___lambda__2(lean_object*); +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__15___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleCompletion___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__22(lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__1; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_FileWorker_handleSemanticTokens_go___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__34___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__14(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_FileWorker_handleSemanticTokens_go___spec__3(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__39(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__30___lambda__1(lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__17___lambda__4___closed__1; -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__25; +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__17___lambda__3___closed__2; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_getInteractiveTermGoal___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleWaitForDiagnostics___lambda__1___closed__1; +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__7___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_NamespaceEntry_finish___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__27___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at_Lean_Server_FileWorker_keywordSemanticTokenMap___spec__2___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__26; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__17(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_11383____spec__23(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__4___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_panic___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__2(lean_object*); 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*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__4___closed__1; 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_Lean_Name_toString(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokens_highlightId(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Server_FileWorker_handleCompletion___lambda__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__26(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getId(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__30(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokens_highlightId___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_RBNode_isRed___rarg(lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__5; @@ -109,139 +102,143 @@ LEAN_EXPORT uint8_t l_Lean_Server_FileWorker_handleHover___lambda__3(lean_object static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__4___closed__2; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleHover___lambda__3___boxed(lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__21___lambda__3___closed__1; static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__2; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__37___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__16___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__7; static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___closed__1; static lean_object* l_Lean_Server_FileWorker_handleCompletion___closed__6; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__30___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Server_FileWorker_getInteractiveGoals___lambda__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__10___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___closed__4; -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__8; uint8_t lean_usize_dec_eq(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokens_highlightKeyword___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__8; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Server_FileWorker_locationLinksOfInfo___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__8; lean_object* l_Lean_Syntax_getArgs(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__9___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleFoldingRange_isImport___closed__2; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_String_Range_includes(lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__2; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__2(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__4___boxed(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_11385____spec__15(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handlePlainGoal___closed__2; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__27___lambda__2(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_11385____spec__23(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__21___lambda__3___closed__1; lean_object* l_Lean_Syntax_getPos_x3f(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDefinition___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__17; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13___lambda__1(lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__30___lambda__3___closed__1; lean_object* l_Lean_Syntax_getTailPos_x3f(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Server_FileWorker_handleSemanticTokens_highlightKeyword___spec__1___boxed(lean_object*, lean_object*); lean_object* l_Lean_Server_RequestM_bindTask___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleCompletion___closed__4; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Server_FileWorker_locationLinksOfInfo___spec__4(lean_object*, lean_object*, size_t, size_t); lean_object* l_Lean_Server_RequestM_withWaitFindSnap___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__2; +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__28(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokensRange(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__29(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__34___lambda__3___closed__1; -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__4; static lean_object* l_Lean_Server_FileWorker_handleCompletion___closed__3; lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9365_(lean_object*); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__6(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____lambda__1(lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__6; +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__16; lean_object* l_Lean_Widget_InteractiveTermGoal_pretty(lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__10; static lean_object* l_Lean_Server_FileWorker_keywordSemanticTokenMap___closed__1; uint8_t l_Char_isAlpha(uint32_t); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addRanges_popRanges___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__19; static lean_object* l_Lean_Server_FileWorker_handlePlainTermGoal___closed__3; -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__5(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__27___lambda__1(lean_object*); static lean_object* l_Lean_Server_FileWorker_NamespaceEntry_finish___closed__2; -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__27___lambda__3___closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__24___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_getInteractiveGoals___lambda__1___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__12; LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Server_FileWorker_getInteractiveGoals___spec__2___lambda__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__9; +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13___lambda__4___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_FileWorker_handleDocumentHighlight___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__29___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__27___lambda__3(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_11385____spec__19___boxed(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__21(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1804_(lean_object*); extern lean_object* l_Lean_Server_requestHandlers; +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__8(lean_object*); static lean_object* l_Lean_Server_FileWorker_NamespaceEntry_finish___closed__5; static lean_object* l_Lean_Server_FileWorker_handleFoldingRange_isImport___closed__1; LEAN_EXPORT uint8_t l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__9(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__10___closed__1; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__27___lambda__3(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_11383____spec__9(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleCompletionItemResolve(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_string_dec_eq(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__30___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__2; -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__6___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_getInteractiveTermGoal___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__17___lambda__3___closed__1; +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__10; static lean_object* l_Lean_Server_FileWorker_handleSemanticTokens_go___closed__2; static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__8; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__4___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols_popStack(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__24___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__19; +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__23___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_string_utf8_byte_size(lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__4; LEAN_EXPORT lean_object* l_List_foldr___at_Lean_Server_FileWorker_NamespaceEntry_finish___spec__1___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__3___closed__1; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__7(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__30(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1509_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__30___lambda__3___closed__2; static lean_object* l_List_mapM_loop___at_Lean_Server_FileWorker_getInteractiveGoals___spec__2___closed__3; +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__24___lambda__2___closed__2; static lean_object* l_List_mapM_loop___at_Lean_Server_FileWorker_getInteractiveGoals___spec__2___closed__4; extern lean_object* l_Lean_projectionFnInfoExt; -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__4___lambda__2___closed__1; lean_object* l_Lean_initializing(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__7___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__10___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleCompletion___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__34___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__30___lambda__3___closed__2; lean_object* l_Lean_Elab_Info_fmtHover_x3f(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handlePlainTermGoal___lambda__1(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__35(lean_object*); static lean_object* l_Lean_Server_FileWorker_handleSemanticTokensFull___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__7___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokensFull___rarg(lean_object*, lean_object*); lean_object* l_Lean_Server_findModuleRefs(lean_object*, lean_object*, uint8_t, uint8_t); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13___lambda__3___closed__2; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__4___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____lambda__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__7___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__1(lean_object*); lean_object* l_Lean_Widget_goalToInteractive(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalContext_pop(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addRanges___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__9; +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__17___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__31(lean_object*); lean_object* l_List_take___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_FileWorker_locationLinksOfInfo___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__37___lambda__3___closed__2; +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__7___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleCompletion___lambda__1___boxed(lean_object*, lean_object*); lean_object* l_Lean_Syntax_isIdOrAtom_x3f(lean_object*); static lean_object* l_List_forIn_loop___at_Lean_Server_FileWorker_handleSemanticTokens___spec__1___lambda__1___closed__1; +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__24___lambda__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__24___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_task_pure(lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__9; +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___closed__1; lean_object* l_Lean_FileMap_lspPosToUtf8Pos(lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__23; LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Server_FileWorker_getInteractiveGoals___spec__2___boxed__const__1; -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__3___closed__2; lean_object* l_Lean_Lsp_ModuleRefs_findAt(lean_object*, lean_object*, uint8_t); static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__5___closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__10___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__4; static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__8; static lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__5; lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainTermGoalParams____x40_Lean_Data_Lsp_Extra___hyg_1307_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__10___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__38(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handlePlainTermGoal(lean_object*, lean_object*, lean_object*); lean_object* lean_local_ctx_find(lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___closed__1; @@ -249,82 +246,80 @@ static lean_object* l_Lean_Server_FileWorker_handlePlainGoal___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); static lean_object* l_List_mapM_loop___at_Lean_Server_FileWorker_getInteractiveGoals___spec__2___closed__5; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__37___lambda__1(lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__34___lambda__3___closed__2; uint8_t lean_expr_eqv(lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleSemanticTokens_addToken___closed__1; -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__27; +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__34___lambda__3___closed__1; static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__8___closed__3; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__37(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__4___lambda__1(lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__10___lambda__4___closed__1; -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13___lambda__4___closed__1; LEAN_EXPORT uint8_t l_Lean_Server_FileWorker_handleHover___lambda__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__21___lambda__3___closed__2; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__3___closed__1; lean_object* l_Lean_RBNode_setBlack___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange___boxed(lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___closed__3; lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3697_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__37___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__37___lambda__3___closed__2; +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__4___lambda__3___closed__1; +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__32(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addCommandRange(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__21___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Server_FileWorker_getInteractiveGoals___spec__2___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__24___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__13; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__34___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13___lambda__3___closed__1; static lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__1; static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__4; static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__8___closed__1; -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__34___lambda__3___closed__2; +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__29; +static lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__2; static lean_object* l_Lean_Server_FileWorker_handleSemanticTokens_go___closed__5; static lean_object* l_Lean_Server_FileWorker_handleCompletion___closed__7; +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__29___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__6; -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__34___lambda__4___closed__1; -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___closed__1; -static lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__5(lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__21___lambda__3___closed__2; +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__2; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__7(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__4; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__8___closed__6; lean_object* l_Lean_Syntax_getKind(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__27___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__9; lean_object* l_Except_map___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_Server_FileWorker_locationLinksOfInfo___spec__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__3___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__21(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addRanges___closed__2; lean_object* l_Lean_Server_Snapshots_Snapshot_env(lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_handlePlainGoal___spec__2___closed__2; -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__11(lean_object*); lean_object* l_Lean_Widget_InteractiveGoal_pretty(lean_object*); lean_object* l_Option_map___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__12; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__7___lambda__1(lean_object*); extern lean_object* l_Lean_Server_RequestError_fileChanged; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol___boxed(lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__18; +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__21___lambda__4___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonPlainGoal____x40_Lean_Data_Lsp_Extra___hyg_1250_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__10___lambda__2(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__10(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__8___closed__4; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addRange___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__21___lambda__4___closed__1; +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__4___lambda__2___closed__2; lean_object* l_IO_sleep(uint32_t, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__4___lambda__3___closed__1; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__37___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__17___lambda__1(lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__19___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__19; +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__10___lambda__4___closed__1; LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Server_FileWorker_handleSemanticTokens_highlightId___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13___lambda__1(lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__3; -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__16; +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__20; LEAN_EXPORT lean_object* l_Array_contains___at_Lean_Server_FileWorker_handleSemanticTokens_go___spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__8___closed__2; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__20___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Server_FileWorker_getInteractiveGoals___spec__3(lean_object*, lean_object*); lean_object* l_Lean_Syntax_getOptional_x3f(lean_object*); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__36___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__6___closed__1; static lean_object* l_List_findSomeM_x3f___at_Lean_Server_FileWorker_handleHover___spec__1___closed__1; lean_object* lean_array_to_list(lean_object*, lean_object*); @@ -332,17 +327,16 @@ LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleWaitForDiagnostics___lam 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*); uint8_t l___private_Lean_Server_GoTo_0__Lean_Server_beqGoToKind____x40_Lean_Server_GoTo___hyg_8_(uint8_t, uint8_t); +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__35(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_getInteractiveGoals___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_keywordSemanticTokenMap; lean_object* l_Lean_FileMap_utf8PosToLspPos(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__15(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_append(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__14(lean_object*); lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoalParams____x40_Lean_Data_Lsp_Extra___hyg_984_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__34___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__12(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___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*); extern lean_object* l_Lean_levelZero; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -350,135 +344,137 @@ LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol(lean_obje lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Server_registerLspRequestHandler___spec__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleHover___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleHover___lambda__7___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__24___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__17___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight___lambda__2(lean_object*, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_Server_FileWorker_locationLinksOfInfo___spec__2___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__1; lean_object* l_Lean_Server_RequestM_asTask___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleHover___lambda__7(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_FileWorker_handleDocumentHighlight___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_constName_x3f(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__21___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_contains___at_Lean_Server_FileWorker_handleSemanticTokens_go___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange(lean_object*); LEAN_EXPORT lean_object* l_List_foldr___at_Lean_Server_FileWorker_NamespaceEntry_finish___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385_(lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Server_FileWorker_handleSemanticTokens___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__3; static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__1; lean_object* l_List_lengthTRAux___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383_(lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__7; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__21___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Task_Priority_default; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__10___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_componentsRev(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleWaitForDiagnostics___lambda__2(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__32(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Lsp_SemanticTokenType_toCtorIdx(uint8_t); lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3643_(lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___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_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__6(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__17___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonDocumentSymbol_go___spec__3(size_t, size_t, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__37___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokens___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__30___lambda__2(lean_object*); +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__39(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MapDeclarationExtension_find_x3f___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_getInteractiveGoals___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Server_locationLinksFromDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__17; static lean_object* l_Lean_Server_FileWorker_handlePlainTermGoal___closed__2; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__10(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__4___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__18; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__21___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__21___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleCompletion___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokens_addToken___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2(lean_object*); lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__29; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__37___lambda__4(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_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2(lean_object*); uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleSemanticTokens___lambda__1___closed__1; LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Server_FileWorker_handleCompletion___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__2; +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__11(lean_object*); static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__8___closed__10; static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___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*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__6; LEAN_EXPORT uint8_t l_Lean_Server_FileWorker_handleHover___lambda__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__12___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__8___closed__11; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__4___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__27___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokens_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_handlePlainGoal___spec__2___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__2___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleWaitForDiagnostics_waitLoop___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addCommandRange___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleWaitForDiagnostics___lambda__1(lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__20; uint8_t l_Lean_LocalDecl_isAuxDecl(lean_object*); static lean_object* l_Lean_Server_FileWorker_handleFoldingRange_isImport___closed__3; static lean_object* l_Lean_Server_FileWorker_handleWaitForDiagnostics___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleWaitForDiagnostics_waitLoop(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handlePlainGoal___lambda__1___closed__3; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__34___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__14; static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___closed__2; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__37___lambda__2(lean_object*); extern lean_object* l_Std_Format_defWidth; +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__21; LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at_Lean_Server_FileWorker_keywordSemanticTokenMap___spec__1(lean_object*, lean_object*, uint8_t); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__12; +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__12(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__4; +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__3___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addRangeFromSyntax___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__8(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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__37___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__6___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__6; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__34___lambda__1(lean_object*); lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonPlainTermGoal____x40_Lean_Data_Lsp_Extra___hyg_1572_(lean_object*); static lean_object* l_Lean_Server_FileWorker_handleSemanticTokens_go___closed__4; -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__28(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleWaitForDiagnostics(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_instInhabitedProjectionFunctionInfo; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__7___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__33___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_String_Range_toLspRange(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f(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_11385____spec__19(lean_object*, lean_object*, lean_object*); lean_object* l_Array_append___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleCompletion(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at_Lean_Server_FileWorker_locationLinksOfInfo___spec__2(lean_object*, lean_object*); lean_object* l_Lean_Server_ModuleRefs_toLspModuleRefs(lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__34(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__24; 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_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____lambda__1(lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__17___lambda__3___closed__1; +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__9(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__34___lambda__3(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_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_handleDocumentHighlight___spec__1(size_t, size_t, lean_object*); lean_object* l_String_csize(uint32_t); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__2; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__16___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__30___lambda__4___closed__1; lean_object* l_IO_AsyncList_waitUntil___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__20(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_List_findSomeM_x3f___at_Lean_Server_FileWorker_handleHover___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Server_FileWorker_handleSemanticTokens_go___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__15; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__7___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_reprint(lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__10___lambda__3___closed__1; lean_object* l_Lean_Server_RequestM_readDoc___at_Lean_Server_RequestM_withWaitFindSnapAtPos___spec__1(lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__10; lean_object* l_panic___at_Lean_Parser_ParserState_mkUnexpectedTokenErrors___spec__1(lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__4___lambda__2___closed__2; lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9158_(lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__13; static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__2___closed__3; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__10___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_handleWaitForDiagnostics___spec__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleCompletion___closed__5; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__17___lambda__1(lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__2; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__34___lambda__2(lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__17___lambda__3___closed__2; +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__9___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__15; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__34___lambda__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange_isImport___boxed(lean_object*); lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3441_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -486,112 +482,118 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_hand static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__8___closed__8; static lean_object* l_Lean_Server_FileWorker_NamespaceEntry_finish___closed__6; uint8_t l_String_Range_contains(lean_object*, lean_object*, uint8_t); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__10___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDefinition___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__39___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addRanges_popRanges(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__9; static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_noHighlightKinds; LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Server_FileWorker_handleSemanticTokens_highlightKeyword___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13___lambda__3___closed__2; static lean_object* l_Lean_Server_FileWorker_NamespaceEntry_finish___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_FileWorker_locationLinksOfInfo___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_string_length(lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Server_FileWorker_handleSemanticTokens_highlightId___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_IO_AsyncList_waitAll___rarg___lambda__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_getInteractiveTermGoal___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__10___lambda__3___closed__2; uint8_t l_Lean_Syntax_hasArgs(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addRangeFromSyntax(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__4___lambda__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__4(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_drop___rarg(lean_object*, lean_object*); lean_object* l_Lean_HashMapImp_find_x3f___at_Lean_Server_References_updateWorkerRefs___spec__3(lean_object*, lean_object*); lean_object* l_Lean_Syntax_findStack_x3f(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__28; LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Server_FileWorker_handleDocumentSymbol___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleHover___closed__2; +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__7; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__7___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__22; static lean_object* l_Lean_Server_FileWorker_handleWaitForDiagnostics___closed__1; static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__8___closed__5; -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__26(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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__4___closed__1; uint8_t lean_uint32_dec_eq(uint32_t, uint32_t); lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonWaitForDiagnosticsParams____x40_Lean_Data_Lsp_Extra___hyg_345_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokens(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Server_FileWorker_locationLinksOfInfo___spec__5(lean_object*, lean_object*, size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__14; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__33(size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__21___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__17___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleCompletionItemResolve___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_findSomeM_x3f___at_Lean_Server_FileWorker_handleHover___spec__1___closed__2; -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__37___lambda__4___closed__1; +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__3___closed__2; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__34___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__25(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addRanges(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__17___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__37(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__30___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances(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_11383____spec__3(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__32___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__34(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__5; static lean_object* l_Lean_Server_FileWorker_NamespaceEntry_finish___closed__4; -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__18; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__10___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__7; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__37___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__22(lean_object*); lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1268_(lean_object*); LEAN_EXPORT uint8_t l_Lean_Server_FileWorker_handleFoldingRange_isImport(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__34___lambda__1(lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__11; +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__39___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDefinition___closed__1; -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__38(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__24; static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__10; uint8_t l_Lean_Syntax_isNone(lean_object*); lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2070_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_getInteractiveGoals___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_span_loop___at_Lean_Server_FileWorker_handleFoldingRange_addRanges___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__24(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__5; lean_object* lean_nat_shiftl(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__16(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Server_FileWorker_getInteractiveGoals___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__21___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_task_map(lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokensFull(lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols_popStack___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__17___lambda__2(lean_object*); lean_object* l_Lean_Server_Completion_find_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_FileWorker_locationLinksOfInfo___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* l_Lean_Expr_getAppFn(lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__37___lambda__4___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__30___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handlePlainTermGoal___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_handleDocumentHighlight___spec__4___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__24(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__3___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__10___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__10(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getRange_x3f(lean_object*, uint8_t); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___closed__7; +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__9; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__5(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___closed__5; static lean_object* l_Lean_Server_FileWorker_handlePlainGoal___lambda__1___closed__2; lean_object* l_Lean_Elab_Info_range_x3f(lean_object*); LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Server_FileWorker_getInteractiveGoals___spec__2___lambda__2(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_mapM_loop___at_Lean_Server_FileWorker_getInteractiveGoals___spec__2___closed__7; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__7___lambda__2(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_Server_FileWorker_locationLinksOfInfo___lambda__8___closed__7; lean_object* l_Lean_Lsp_CompletionItem_getFileSource_x21(lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__17___lambda__5(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___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__17___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__27___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__30___lambda__4___closed__1; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__21___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__37___lambda__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handlePlainGoal(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleSemanticTokens_go___closed__3; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__4___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_handleWaitForDiagnostics___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Server_FileWorker_handleSemanticTokens___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__27___lambda__3___closed__1; static lean_object* l_List_findSomeM_x3f___at_Lean_Server_FileWorker_handleHover___spec__1___closed__4; static lean_object* l_Lean_Server_FileWorker_keywordSemanticTokenMap___closed__2; lean_object* l_List_reverse___rarg(lean_object*); @@ -599,33 +601,29 @@ LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleCompletionItemResolve___ LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Server_FileWorker_handleSemanticTokens_go___spec__2(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_handleDocumentHighlight___spec__2(size_t, size_t, lean_object*); static lean_object* l_List_findSomeM_x3f___at_Lean_Server_FileWorker_handleHover___spec__1___closed__3; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__24___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_String_intercalate(lean_object*, lean_object*); lean_object* l_Lean_Server_Completion_fillEligibleHeaderDecls(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__3(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__32___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__24___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__5; static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__14; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__17(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_pure___at_Lean_Server_RefInfo_toLspRefInfo___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__24___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__28; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_handleDocumentHighlight___spec__4(size_t, size_t, lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__16; -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__12___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokens_highlightKeyword(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___closed__6; -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__37___lambda__3___closed__1; -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__11; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__4___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; static lean_object* l_Lean_Server_FileWorker_keywordSemanticTokenMap___closed__3; lean_object* lean_array_uget(lean_object*, size_t); lean_object* l_Lean_Elab_Info_lctx(lean_object*); LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Server_FileWorker_handleCompletion___spec__1(lean_object*); lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_350_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__10___lambda__2(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_getInteractiveGoals___spec__1___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_Server_FileWorker_locationLinksOfInfo___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_getInteractiveTermGoal___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -633,100 +631,103 @@ lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover_ LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Server_FileWorker_locationLinksOfInfo___spec__3(lean_object*, lean_object*, size_t, size_t); lean_object* lean_io_error_to_string(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__16(size_t, size_t, lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_InfoTree_findInfo_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_getInteractiveTermGoal(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__31(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__18(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleHover___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Info_pos_x3f(lean_object*); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__19(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__24___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__27; lean_object* l_Lean_Server_Completion_resolveCompletionItem_x3f(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_object* l_IO_AsyncList_getFinishedPrefix___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__17___lambda__2(lean_object*); lean_object* l_List_redLength___rarg(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__33___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleSemanticTokens_go___closed__6; lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__7; -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__23___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__4(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__36(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__21___lambda__4(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_Server_FileWorker_handleFoldingRange_addCommandRange___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Server_FileWorker_handleCompletion___spec__1___rarg(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_PersistentHashMap_contains___at_Lean_Server_registerLspRequestHandler___spec__6(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__21___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__10___lambda__3___closed__2; +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__1; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokensFull___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__4___lambda__1(lean_object*); lean_object* lean_array_get_size(lean_object*); static lean_object* l_Lean_Server_FileWorker_handlePlainGoal___closed__3; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__7___lambda__2___closed__2; +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__23; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__20(size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__10___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__30___lambda__1(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__21___lambda__2(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__27___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_FileWorker_handleDocumentHighlight___spec__3(size_t, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__17___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__15; lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Server_FileWorker_getInteractiveGoals___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_NamespaceEntry_finish___closed__3; +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__34___lambda__4___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokens_addToken(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__36(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__6; LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Server_FileWorker_getInteractiveGoals___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); static lean_object* l_List_mapM_loop___at_Lean_Server_FileWorker_getInteractiveGoals___spec__2___closed__6; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleCompletion___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__27(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2013_(lean_object*); static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__2___closed__2; static lean_object* l_Lean_Server_FileWorker_handleFoldingRange_isImport___closed__5; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_handlePlainGoal___spec__2(size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____lambda__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__20___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokens_highlightId___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at_Lean_Server_FileWorker_keywordSemanticTokenMap___spec__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__3; lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2387_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__2(lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__24___lambda__4(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_11385____spec__36___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__24___lambda__3___closed__1; lean_object* l_Lean_Elab_InfoTree_hoverableInfoAt_x3f(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t); lean_object* l_Lean_Elab_Info_stx(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13___lambda__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13___lambda__2(lean_object*); uint8_t l_Lean_Exception_isRuntime(lean_object*); lean_object* l_Lean_Server_RequestError_ofIoError(lean_object*); uint8_t lean_string_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Server_FileWorker_handleSemanticTokens_highlightId___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__30; static lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addRanges___closed__1; static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__8___closed__9; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__34___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Info_toElabInfo_x3f(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__8(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleHover___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleFoldingRange_isImport___closed__4; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__37___lambda__1(lean_object*); lean_object* l_Lean_Expr_mvarId_x21(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__30___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__7___lambda__2___closed__1; static lean_object* l_List_mapM_loop___at_Lean_Server_FileWorker_getInteractiveGoals___spec__2___closed__1; static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__7; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__7___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__30___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___lambda__1(lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__4___closed__2; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addRange(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__3; static lean_object* l_Lean_Server_FileWorker_handleHover___closed__1; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__7___lambda__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__24___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__3; -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__17; lean_object* l_Lean_Elab_ContextInfo_runMetaM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5066_(lean_object*); lean_object* l_Lean_Server_Snapshots_Snapshot_endPos(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__4(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__30___lambda__4___boxed(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_11383____spec__9___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__15___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__2___boxed(lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__25; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__37___lambda__4___boxed(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_11385____spec__3(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__29(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__27(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__22; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleWaitForDiagnostics___lambda__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Server_FileWorker_handleSemanticTokens___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at_Lean_Server_FileWorker_keywordSemanticTokenMap___spec__2(lean_object*, lean_object*, uint8_t); @@ -734,7 +735,6 @@ LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Server_FileWorker_handleDocument LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol___rarg(lean_object*, lean_object*); lean_object* l_List_toArrayAux___rarg(lean_object*, lean_object*); lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8778_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_findSomeM_x3f___at_Lean_Server_FileWorker_handleHover___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDefinition___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_getInteractiveGoals(lean_object*, lean_object*, lean_object*); @@ -22133,7 +22133,7 @@ lean_dec(x_1); return x_4; } } -static lean_object* _init_l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__1() { +static lean_object* _init_l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__1() { _start: { lean_object* x_1; @@ -22141,7 +22141,7 @@ x_1 = lean_mk_string_from_bytes("Cannot parse request params: ", 29); return x_1; } } -static lean_object* _init_l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__2() { +static lean_object* _init_l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__2() { _start: { lean_object* x_1; @@ -22149,7 +22149,7 @@ x_1 = lean_mk_string_from_bytes("\n", 1); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2(lean_object* x_1) { _start: { lean_object* x_2; @@ -22163,10 +22163,10 @@ if (x_3 == 0) lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; x_4 = lean_ctor_get(x_2, 0); x_5 = l_Lean_Json_compress(x_1); -x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__1; +x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__1; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); -x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__2; +x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__2; x_9 = lean_string_append(x_7, x_8); x_10 = lean_string_append(x_9, x_4); lean_dec(x_4); @@ -22186,10 +22186,10 @@ x_15 = lean_ctor_get(x_2, 0); lean_inc(x_15); lean_dec(x_2); x_16 = l_Lean_Json_compress(x_1); -x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__1; +x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__1; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); -x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__2; +x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__2; x_20 = lean_string_append(x_18, x_19); x_21 = lean_string_append(x_20, x_15); lean_dec(x_15); @@ -22226,7 +22226,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -22251,11 +22251,11 @@ return x_7; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2(x_1); +x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -22305,7 +22305,7 @@ return x_11; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__2___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__2___closed__1() { _start: { lean_object* x_1; lean_object* x_2; @@ -22314,38 +22314,38 @@ x_2 = l_Lean_Json_mkObj(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__2(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__2___closed__1; +x_2 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__2___closed__1; return x_2; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__3___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__3___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__2___boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__2___boxed), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__3___closed__2() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__3___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__3___closed__1; +x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__3___closed__1; x_2 = lean_alloc_closure((void*)(l_Except_map___rarg), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; -x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2(x_2); -x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__3(x_5, x_3, x_4); +x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2(x_2); +x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__3(x_5, x_3, x_4); lean_dec(x_5); if (lean_obj_tag(x_6) == 0) { @@ -22364,7 +22364,7 @@ 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 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__3___closed__2; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__3___closed__2; x_13 = l_Task_Priority_default; x_14 = 0; x_15 = lean_task_map(x_12, x_11, x_13, x_14); @@ -22379,7 +22379,7 @@ x_17 = lean_ctor_get(x_9, 1); lean_inc(x_17); lean_inc(x_16); lean_dec(x_9); -x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__3___closed__2; +x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__3___closed__2; x_19 = l_Task_Priority_default; x_20 = 0; x_21 = lean_task_map(x_18, x_16, x_19, x_20); @@ -22438,7 +22438,7 @@ return x_30; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__4___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__4___closed__1() { _start: { lean_object* x_1; @@ -22446,28 +22446,28 @@ x_1 = l_Lean_Server_requestHandlers; return x_1; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__4___closed__2() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__4___closed__2() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__1), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____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_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___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_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_6 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__3), 4, 1); +x_6 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__3), 4, 1); lean_closure_set(x_6, 0, x_1); -x_7 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__4___closed__1; +x_7 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__4___closed__1; x_8 = lean_st_ref_take(x_7, x_5); 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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__4___closed__2; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__4___closed__2; x_12 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_6); @@ -22494,7 +22494,7 @@ return x_18; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1() { _start: { lean_object* x_1; @@ -22502,7 +22502,7 @@ x_1 = lean_mk_string_from_bytes("Failed to register LSP request handler for '", return x_1; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__2() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__2() { _start: { lean_object* x_1; @@ -22510,12 +22510,12 @@ x_1 = lean_mk_string_from_bytes("': already registered", 21); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__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; lean_object* x_7; uint8_t x_8; lean_dec(x_4); -x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__4___closed__1; +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__4___closed__1; x_7 = lean_st_ref_get(x_6, x_5); x_8 = !lean_is_exclusive(x_7); if (x_8 == 0) @@ -22530,7 +22530,7 @@ if (x_11 == 0) lean_object* x_12; lean_object* x_13; lean_free_object(x_7); x_12 = lean_box(0); -x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__4(x_1, x_2, x_3, x_12, x_10); +x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__4(x_1, x_2, x_3, x_12, x_10); return x_13; } else @@ -22538,10 +22538,10 @@ else lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_dec(x_2); lean_dec(x_1); -x_14 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_14 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_15 = lean_string_append(x_14, x_3); lean_dec(x_3); -x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__2; +x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__2; x_17 = lean_string_append(x_15, x_16); x_18 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_18, 0, x_17); @@ -22564,7 +22564,7 @@ if (x_21 == 0) { lean_object* x_22; lean_object* x_23; x_22 = lean_box(0); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__4(x_1, x_2, x_3, x_22, x_20); +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__4(x_1, x_2, x_3, x_22, x_20); return x_23; } else @@ -22572,10 +22572,10 @@ 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_dec(x_2); lean_dec(x_1); -x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_25 = lean_string_append(x_24, x_3); lean_dec(x_3); -x_26 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__2; +x_26 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__2; x_27 = lean_string_append(x_25, x_26); x_28 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_28, 0, x_27); @@ -22587,7 +22587,7 @@ return x_29; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___closed__1() { _start: { lean_object* x_1; @@ -22595,7 +22595,7 @@ x_1 = lean_mk_string_from_bytes("': only possible during initialization", 38); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____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; @@ -22615,10 +22615,10 @@ if (x_8 == 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; x_9 = lean_ctor_get(x_5, 0); lean_dec(x_9); -x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_11 = lean_string_append(x_10, x_1); lean_dec(x_1); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___closed__1; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___closed__1; x_13 = lean_string_append(x_11, x_12); x_14 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_14, 0, x_13); @@ -22632,10 +22632,10 @@ lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean x_15 = lean_ctor_get(x_5, 1); lean_inc(x_15); lean_dec(x_5); -x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_17 = lean_string_append(x_16, x_1); lean_dec(x_1); -x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___closed__1; +x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___closed__1; x_19 = lean_string_append(x_17, x_18); x_20 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_20, 0, x_19); @@ -22652,12 +22652,12 @@ x_22 = lean_ctor_get(x_5, 1); lean_inc(x_22); lean_dec(x_5); x_23 = lean_box(0); -x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5(x_2, x_3, x_1, x_23, x_22); +x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5(x_2, x_3, x_1, x_23, x_22); return x_24; } } } -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__5(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__5(lean_object* x_1) { _start: { lean_object* x_2; @@ -22671,10 +22671,10 @@ if (x_3 == 0) lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; x_4 = lean_ctor_get(x_2, 0); x_5 = l_Lean_Json_compress(x_1); -x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__1; +x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__1; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); -x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__2; +x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__2; x_9 = lean_string_append(x_7, x_8); x_10 = lean_string_append(x_9, x_4); lean_dec(x_4); @@ -22694,10 +22694,10 @@ x_15 = lean_ctor_get(x_2, 0); lean_inc(x_15); lean_dec(x_2); x_16 = l_Lean_Json_compress(x_1); -x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__1; +x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__1; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); -x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__2; +x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__2; x_20 = lean_string_append(x_18, x_19); x_21 = lean_string_append(x_20, x_15); lean_dec(x_15); @@ -22734,7 +22734,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -22759,11 +22759,11 @@ return x_7; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__4___lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__4___lambda__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__5(x_1); +x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__5(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -22813,7 +22813,7 @@ return x_11; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__4___lambda__2___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__4___lambda__2___closed__1() { _start: { lean_object* x_1; @@ -22821,22 +22821,22 @@ x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__L return x_1; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__4___lambda__2___closed__2() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__4___lambda__2___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__4___lambda__2___closed__1; +x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__4___lambda__2___closed__1; x_2 = lean_alloc_closure((void*)(l_Except_map___rarg), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__4___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__4___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; -x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__5(x_2); -x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__6(x_5, x_3, x_4); +x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__5(x_2); +x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__6(x_5, x_3, x_4); lean_dec(x_5); if (lean_obj_tag(x_6) == 0) { @@ -22855,7 +22855,7 @@ 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 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__4___lambda__2___closed__2; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__4___lambda__2___closed__2; x_13 = l_Task_Priority_default; x_14 = 0; x_15 = lean_task_map(x_12, x_11, x_13, x_14); @@ -22870,7 +22870,7 @@ x_17 = lean_ctor_get(x_9, 1); lean_inc(x_17); lean_inc(x_16); lean_dec(x_9); -x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__4___lambda__2___closed__2; +x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__4___lambda__2___closed__2; x_19 = l_Task_Priority_default; x_20 = 0; x_21 = lean_task_map(x_18, x_16, x_19, x_20); @@ -22929,28 +22929,28 @@ return x_30; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__4___lambda__3___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__4___lambda__3___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__4___lambda__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__4___lambda__1), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__4___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__4___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_6 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__4___lambda__2), 4, 1); +x_6 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__4___lambda__2), 4, 1); lean_closure_set(x_6, 0, x_1); -x_7 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__4___closed__1; +x_7 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__4___closed__1; x_8 = lean_st_ref_take(x_7, x_5); 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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__4___lambda__3___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__4___lambda__3___closed__1; x_12 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_6); @@ -22977,12 +22977,12 @@ return x_18; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__4___lambda__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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__4___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_6; lean_object* x_7; uint8_t x_8; lean_dec(x_4); -x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__4___closed__1; +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__4___closed__1; x_7 = lean_st_ref_get(x_6, x_5); x_8 = !lean_is_exclusive(x_7); if (x_8 == 0) @@ -22997,7 +22997,7 @@ if (x_11 == 0) lean_object* x_12; lean_object* x_13; lean_free_object(x_7); x_12 = lean_box(0); -x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__4___lambda__3(x_1, x_2, x_3, x_12, x_10); +x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__4___lambda__3(x_1, x_2, x_3, x_12, x_10); return x_13; } else @@ -23005,10 +23005,10 @@ else lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_dec(x_2); lean_dec(x_1); -x_14 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_14 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_15 = lean_string_append(x_14, x_3); lean_dec(x_3); -x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__2; +x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__2; x_17 = lean_string_append(x_15, x_16); x_18 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_18, 0, x_17); @@ -23031,7 +23031,7 @@ if (x_21 == 0) { lean_object* x_22; lean_object* x_23; x_22 = lean_box(0); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__4___lambda__3(x_1, x_2, x_3, x_22, x_20); +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__4___lambda__3(x_1, x_2, x_3, x_22, x_20); return x_23; } else @@ -23039,10 +23039,10 @@ 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_dec(x_2); lean_dec(x_1); -x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_25 = lean_string_append(x_24, x_3); lean_dec(x_3); -x_26 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__2; +x_26 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__2; x_27 = lean_string_append(x_25, x_26); x_28 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_28, 0, x_27); @@ -23054,7 +23054,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; uint8_t x_7; @@ -23074,10 +23074,10 @@ if (x_8 == 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; x_9 = lean_ctor_get(x_5, 0); lean_dec(x_9); -x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_11 = lean_string_append(x_10, x_1); lean_dec(x_1); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___closed__1; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___closed__1; x_13 = lean_string_append(x_11, x_12); x_14 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_14, 0, x_13); @@ -23091,10 +23091,10 @@ lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean x_15 = lean_ctor_get(x_5, 1); lean_inc(x_15); lean_dec(x_5); -x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_17 = lean_string_append(x_16, x_1); lean_dec(x_1); -x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___closed__1; +x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___closed__1; x_19 = lean_string_append(x_17, x_18); x_20 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_20, 0, x_19); @@ -23111,12 +23111,12 @@ x_22 = lean_ctor_get(x_5, 1); lean_inc(x_22); lean_dec(x_5); x_23 = lean_box(0); -x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__4___lambda__4(x_2, x_3, x_1, x_23, x_22); +x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__4___lambda__4(x_2, x_3, x_1, x_23, x_22); return x_24; } } } -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__8(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__8(lean_object* x_1) { _start: { lean_object* x_2; @@ -23130,10 +23130,10 @@ if (x_3 == 0) lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; x_4 = lean_ctor_get(x_2, 0); x_5 = l_Lean_Json_compress(x_1); -x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__1; +x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__1; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); -x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__2; +x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__2; x_9 = lean_string_append(x_7, x_8); x_10 = lean_string_append(x_9, x_4); lean_dec(x_4); @@ -23153,10 +23153,10 @@ x_15 = lean_ctor_get(x_2, 0); lean_inc(x_15); lean_dec(x_2); x_16 = l_Lean_Json_compress(x_1); -x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__1; +x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__1; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); -x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__2; +x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__2; x_20 = lean_string_append(x_18, x_19); x_21 = lean_string_append(x_20, x_15); lean_dec(x_15); @@ -23193,7 +23193,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -23218,11 +23218,11 @@ return x_7; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__7___lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__7___lambda__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__8(x_1); +x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__8(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -23268,7 +23268,7 @@ return x_11; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__7___lambda__2___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__7___lambda__2___closed__1() { _start: { lean_object* x_1; @@ -23276,22 +23276,22 @@ x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__L return x_1; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__7___lambda__2___closed__2() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__7___lambda__2___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__7___lambda__2___closed__1; +x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__7___lambda__2___closed__1; x_2 = lean_alloc_closure((void*)(l_Except_map___rarg), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__7___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__7___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; -x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__8(x_2); -x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__9(x_5, x_3, x_4); +x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__8(x_2); +x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__9(x_5, x_3, x_4); lean_dec(x_5); if (lean_obj_tag(x_6) == 0) { @@ -23310,7 +23310,7 @@ 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 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__7___lambda__2___closed__2; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__7___lambda__2___closed__2; x_13 = l_Task_Priority_default; x_14 = 0; x_15 = lean_task_map(x_12, x_11, x_13, x_14); @@ -23325,7 +23325,7 @@ x_17 = lean_ctor_get(x_9, 1); lean_inc(x_17); lean_inc(x_16); lean_dec(x_9); -x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__7___lambda__2___closed__2; +x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__7___lambda__2___closed__2; x_19 = l_Task_Priority_default; x_20 = 0; x_21 = lean_task_map(x_18, x_16, x_19, x_20); @@ -23384,28 +23384,28 @@ return x_30; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__7___lambda__3___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__7___lambda__3___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__7___lambda__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__7___lambda__1), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__7___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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__7___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_6 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__7___lambda__2), 4, 1); +x_6 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__7___lambda__2), 4, 1); lean_closure_set(x_6, 0, x_1); -x_7 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__4___closed__1; +x_7 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__4___closed__1; x_8 = lean_st_ref_take(x_7, x_5); 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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__7___lambda__3___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__7___lambda__3___closed__1; x_12 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_6); @@ -23432,12 +23432,12 @@ return x_18; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__7___lambda__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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__7___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_6; lean_object* x_7; uint8_t x_8; lean_dec(x_4); -x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__4___closed__1; +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__4___closed__1; x_7 = lean_st_ref_get(x_6, x_5); x_8 = !lean_is_exclusive(x_7); if (x_8 == 0) @@ -23452,7 +23452,7 @@ if (x_11 == 0) lean_object* x_12; lean_object* x_13; lean_free_object(x_7); x_12 = lean_box(0); -x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__7___lambda__3(x_1, x_2, x_3, x_12, x_10); +x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__7___lambda__3(x_1, x_2, x_3, x_12, x_10); return x_13; } else @@ -23460,10 +23460,10 @@ else lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_dec(x_2); lean_dec(x_1); -x_14 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_14 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_15 = lean_string_append(x_14, x_3); lean_dec(x_3); -x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__2; +x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__2; x_17 = lean_string_append(x_15, x_16); x_18 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_18, 0, x_17); @@ -23486,7 +23486,7 @@ if (x_21 == 0) { lean_object* x_22; lean_object* x_23; x_22 = lean_box(0); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__7___lambda__3(x_1, x_2, x_3, x_22, x_20); +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__7___lambda__3(x_1, x_2, x_3, x_22, x_20); return x_23; } else @@ -23494,10 +23494,10 @@ 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_dec(x_2); lean_dec(x_1); -x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_25 = lean_string_append(x_24, x_3); lean_dec(x_3); -x_26 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__2; +x_26 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__2; x_27 = lean_string_append(x_25, x_26); x_28 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_28, 0, x_27); @@ -23509,7 +23509,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____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; uint8_t x_7; @@ -23529,10 +23529,10 @@ if (x_8 == 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; x_9 = lean_ctor_get(x_5, 0); lean_dec(x_9); -x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_11 = lean_string_append(x_10, x_1); lean_dec(x_1); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___closed__1; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___closed__1; x_13 = lean_string_append(x_11, x_12); x_14 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_14, 0, x_13); @@ -23546,10 +23546,10 @@ lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean x_15 = lean_ctor_get(x_5, 1); lean_inc(x_15); lean_dec(x_5); -x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_17 = lean_string_append(x_16, x_1); lean_dec(x_1); -x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___closed__1; +x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___closed__1; x_19 = lean_string_append(x_17, x_18); x_20 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_20, 0, x_19); @@ -23566,12 +23566,12 @@ x_22 = lean_ctor_get(x_5, 1); lean_inc(x_22); lean_dec(x_5); x_23 = lean_box(0); -x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__7___lambda__4(x_2, x_3, x_1, x_23, x_22); +x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__7___lambda__4(x_2, x_3, x_1, x_23, x_22); return x_24; } } } -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__11(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__11(lean_object* x_1) { _start: { lean_object* x_2; @@ -23585,10 +23585,10 @@ if (x_3 == 0) lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; x_4 = lean_ctor_get(x_2, 0); x_5 = l_Lean_Json_compress(x_1); -x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__1; +x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__1; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); -x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__2; +x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__2; x_9 = lean_string_append(x_7, x_8); x_10 = lean_string_append(x_9, x_4); lean_dec(x_4); @@ -23608,10 +23608,10 @@ x_15 = lean_ctor_get(x_2, 0); lean_inc(x_15); lean_dec(x_2); x_16 = l_Lean_Json_compress(x_1); -x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__1; +x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__1; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); -x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__2; +x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__2; x_20 = lean_string_append(x_18, x_19); x_21 = lean_string_append(x_20, x_15); lean_dec(x_15); @@ -23648,7 +23648,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -23673,11 +23673,11 @@ return x_7; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__10___lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__10___lambda__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__11(x_1); +x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__11(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -23727,7 +23727,7 @@ return x_11; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__10___lambda__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__10___lambda__2(lean_object* x_1) { _start: { if (lean_obj_tag(x_1) == 0) @@ -23747,30 +23747,30 @@ return x_4; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__10___lambda__3___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__10___lambda__3___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__10___lambda__2), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__10___lambda__2), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__10___lambda__3___closed__2() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__10___lambda__3___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__10___lambda__3___closed__1; +x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__10___lambda__3___closed__1; x_2 = lean_alloc_closure((void*)(l_Except_map___rarg), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__10___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__10___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; -x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__11(x_2); -x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__12(x_5, x_3, x_4); +x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__11(x_2); +x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__12(x_5, x_3, x_4); lean_dec(x_5); if (lean_obj_tag(x_6) == 0) { @@ -23789,7 +23789,7 @@ 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 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__10___lambda__3___closed__2; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__10___lambda__3___closed__2; x_13 = l_Task_Priority_default; x_14 = 0; x_15 = lean_task_map(x_12, x_11, x_13, x_14); @@ -23804,7 +23804,7 @@ x_17 = lean_ctor_get(x_9, 1); lean_inc(x_17); lean_inc(x_16); lean_dec(x_9); -x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__10___lambda__3___closed__2; +x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__10___lambda__3___closed__2; x_19 = l_Task_Priority_default; x_20 = 0; x_21 = lean_task_map(x_18, x_16, x_19, x_20); @@ -23863,28 +23863,28 @@ return x_30; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__10___lambda__4___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__10___lambda__4___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__10___lambda__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__10___lambda__1), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__10___lambda__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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__10___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_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_6 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__10___lambda__3), 4, 1); +x_6 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__10___lambda__3), 4, 1); lean_closure_set(x_6, 0, x_1); -x_7 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__4___closed__1; +x_7 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__4___closed__1; x_8 = lean_st_ref_take(x_7, x_5); 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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__10___lambda__4___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__10___lambda__4___closed__1; x_12 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_6); @@ -23911,12 +23911,12 @@ return x_18; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__10___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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__10___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; lean_object* x_7; uint8_t x_8; lean_dec(x_4); -x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__4___closed__1; +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__4___closed__1; x_7 = lean_st_ref_get(x_6, x_5); x_8 = !lean_is_exclusive(x_7); if (x_8 == 0) @@ -23931,7 +23931,7 @@ if (x_11 == 0) lean_object* x_12; lean_object* x_13; lean_free_object(x_7); x_12 = lean_box(0); -x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__10___lambda__4(x_1, x_2, x_3, x_12, x_10); +x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__10___lambda__4(x_1, x_2, x_3, x_12, x_10); return x_13; } else @@ -23939,10 +23939,10 @@ else lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_dec(x_2); lean_dec(x_1); -x_14 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_14 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_15 = lean_string_append(x_14, x_3); lean_dec(x_3); -x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__2; +x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__2; x_17 = lean_string_append(x_15, x_16); x_18 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_18, 0, x_17); @@ -23965,7 +23965,7 @@ if (x_21 == 0) { lean_object* x_22; lean_object* x_23; x_22 = lean_box(0); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__10___lambda__4(x_1, x_2, x_3, x_22, x_20); +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__10___lambda__4(x_1, x_2, x_3, x_22, x_20); return x_23; } else @@ -23973,10 +23973,10 @@ 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_dec(x_2); lean_dec(x_1); -x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_25 = lean_string_append(x_24, x_3); lean_dec(x_3); -x_26 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__2; +x_26 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__2; x_27 = lean_string_append(x_25, x_26); x_28 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_28, 0, x_27); @@ -23988,7 +23988,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____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; uint8_t x_7; @@ -24008,10 +24008,10 @@ if (x_8 == 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; x_9 = lean_ctor_get(x_5, 0); lean_dec(x_9); -x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_11 = lean_string_append(x_10, x_1); lean_dec(x_1); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___closed__1; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___closed__1; x_13 = lean_string_append(x_11, x_12); x_14 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_14, 0, x_13); @@ -24025,10 +24025,10 @@ lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean x_15 = lean_ctor_get(x_5, 1); lean_inc(x_15); lean_dec(x_5); -x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_17 = lean_string_append(x_16, x_1); lean_dec(x_1); -x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___closed__1; +x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___closed__1; x_19 = lean_string_append(x_17, x_18); x_20 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_20, 0, x_19); @@ -24045,12 +24045,12 @@ x_22 = lean_ctor_get(x_5, 1); lean_inc(x_22); lean_dec(x_5); x_23 = lean_box(0); -x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__10___lambda__5(x_2, x_3, x_1, x_23, x_22); +x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__10___lambda__5(x_2, x_3, x_1, x_23, x_22); return x_24; } } } -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__14(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__14(lean_object* x_1) { _start: { lean_object* x_2; @@ -24064,10 +24064,10 @@ if (x_3 == 0) lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; x_4 = lean_ctor_get(x_2, 0); x_5 = l_Lean_Json_compress(x_1); -x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__1; +x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__1; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); -x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__2; +x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__2; x_9 = lean_string_append(x_7, x_8); x_10 = lean_string_append(x_9, x_4); lean_dec(x_4); @@ -24087,10 +24087,10 @@ x_15 = lean_ctor_get(x_2, 0); lean_inc(x_15); lean_dec(x_2); x_16 = l_Lean_Json_compress(x_1); -x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__1; +x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__1; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); -x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__2; +x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__2; x_20 = lean_string_append(x_18, x_19); x_21 = lean_string_append(x_20, x_15); lean_dec(x_15); @@ -24127,7 +24127,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__15(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__15(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -24152,7 +24152,7 @@ return x_7; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__16(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__16(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -24177,11 +24177,11 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13___lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13___lambda__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__14(x_1); +x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__14(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -24231,7 +24231,7 @@ return x_11; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13___lambda__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13___lambda__2(lean_object* x_1) { _start: { lean_object* x_2; size_t x_3; size_t x_4; lean_object* x_5; lean_object* x_6; @@ -24239,36 +24239,36 @@ x_2 = lean_array_get_size(x_1); x_3 = lean_usize_of_nat(x_2); lean_dec(x_2); x_4 = 0; -x_5 = l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__16(x_3, x_4, x_1); +x_5 = l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__16(x_3, x_4, x_1); x_6 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_6, 0, x_5); return x_6; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13___lambda__3___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13___lambda__3___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13___lambda__2), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13___lambda__2), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13___lambda__3___closed__2() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13___lambda__3___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13___lambda__3___closed__1; +x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13___lambda__3___closed__1; x_2 = lean_alloc_closure((void*)(l_Except_map___rarg), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; -x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__14(x_2); -x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__15(x_5, x_3, x_4); +x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__14(x_2); +x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__15(x_5, x_3, x_4); lean_dec(x_5); if (lean_obj_tag(x_6) == 0) { @@ -24287,7 +24287,7 @@ 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 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13___lambda__3___closed__2; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13___lambda__3___closed__2; x_13 = l_Task_Priority_default; x_14 = 0; x_15 = lean_task_map(x_12, x_11, x_13, x_14); @@ -24302,7 +24302,7 @@ x_17 = lean_ctor_get(x_9, 1); lean_inc(x_17); lean_inc(x_16); lean_dec(x_9); -x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13___lambda__3___closed__2; +x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13___lambda__3___closed__2; x_19 = l_Task_Priority_default; x_20 = 0; x_21 = lean_task_map(x_18, x_16, x_19, x_20); @@ -24361,28 +24361,28 @@ return x_30; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13___lambda__4___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13___lambda__4___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13___lambda__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13___lambda__1), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13___lambda__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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13___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_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_6 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13___lambda__3), 4, 1); +x_6 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13___lambda__3), 4, 1); lean_closure_set(x_6, 0, x_1); -x_7 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__4___closed__1; +x_7 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__4___closed__1; x_8 = lean_st_ref_take(x_7, x_5); 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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13___lambda__4___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13___lambda__4___closed__1; x_12 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_6); @@ -24409,12 +24409,12 @@ return x_18; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13___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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13___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; lean_object* x_7; uint8_t x_8; lean_dec(x_4); -x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__4___closed__1; +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__4___closed__1; x_7 = lean_st_ref_get(x_6, x_5); x_8 = !lean_is_exclusive(x_7); if (x_8 == 0) @@ -24429,7 +24429,7 @@ if (x_11 == 0) lean_object* x_12; lean_object* x_13; lean_free_object(x_7); x_12 = lean_box(0); -x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13___lambda__4(x_1, x_2, x_3, x_12, x_10); +x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13___lambda__4(x_1, x_2, x_3, x_12, x_10); return x_13; } else @@ -24437,10 +24437,10 @@ else lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_dec(x_2); lean_dec(x_1); -x_14 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_14 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_15 = lean_string_append(x_14, x_3); lean_dec(x_3); -x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__2; +x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__2; x_17 = lean_string_append(x_15, x_16); x_18 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_18, 0, x_17); @@ -24463,7 +24463,7 @@ if (x_21 == 0) { lean_object* x_22; lean_object* x_23; x_22 = lean_box(0); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13___lambda__4(x_1, x_2, x_3, x_22, x_20); +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13___lambda__4(x_1, x_2, x_3, x_22, x_20); return x_23; } else @@ -24471,10 +24471,10 @@ 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_dec(x_2); lean_dec(x_1); -x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_25 = lean_string_append(x_24, x_3); lean_dec(x_3); -x_26 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__2; +x_26 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__2; x_27 = lean_string_append(x_25, x_26); x_28 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_28, 0, x_27); @@ -24486,7 +24486,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____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; uint8_t x_7; @@ -24506,10 +24506,10 @@ if (x_8 == 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; x_9 = lean_ctor_get(x_5, 0); lean_dec(x_9); -x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_11 = lean_string_append(x_10, x_1); lean_dec(x_1); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___closed__1; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___closed__1; x_13 = lean_string_append(x_11, x_12); x_14 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_14, 0, x_13); @@ -24523,10 +24523,10 @@ lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean x_15 = lean_ctor_get(x_5, 1); lean_inc(x_15); lean_dec(x_5); -x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_17 = lean_string_append(x_16, x_1); lean_dec(x_1); -x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___closed__1; +x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___closed__1; x_19 = lean_string_append(x_17, x_18); x_20 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_20, 0, x_19); @@ -24543,12 +24543,12 @@ x_22 = lean_ctor_get(x_5, 1); lean_inc(x_22); lean_dec(x_5); x_23 = lean_box(0); -x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13___lambda__5(x_2, x_3, x_1, x_23, x_22); +x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13___lambda__5(x_2, x_3, x_1, x_23, x_22); return x_24; } } } -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__18(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__18(lean_object* x_1) { _start: { lean_object* x_2; @@ -24562,10 +24562,10 @@ if (x_3 == 0) lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; x_4 = lean_ctor_get(x_2, 0); x_5 = l_Lean_Json_compress(x_1); -x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__1; +x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__1; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); -x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__2; +x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__2; x_9 = lean_string_append(x_7, x_8); x_10 = lean_string_append(x_9, x_4); lean_dec(x_4); @@ -24585,10 +24585,10 @@ x_15 = lean_ctor_get(x_2, 0); lean_inc(x_15); lean_dec(x_2); x_16 = l_Lean_Json_compress(x_1); -x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__1; +x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__1; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); -x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__2; +x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__2; x_20 = lean_string_append(x_18, x_19); x_21 = lean_string_append(x_20, x_15); lean_dec(x_15); @@ -24625,7 +24625,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__19(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__19(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -24650,7 +24650,7 @@ return x_7; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__20(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__20(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -24675,11 +24675,11 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__17___lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__17___lambda__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__18(x_1); +x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__18(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -24729,7 +24729,7 @@ return x_11; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__17___lambda__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__17___lambda__2(lean_object* x_1) { _start: { lean_object* x_2; size_t x_3; size_t x_4; lean_object* x_5; lean_object* x_6; @@ -24737,36 +24737,36 @@ x_2 = lean_array_get_size(x_1); x_3 = lean_usize_of_nat(x_2); lean_dec(x_2); x_4 = 0; -x_5 = l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__20(x_3, x_4, x_1); +x_5 = l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__20(x_3, x_4, x_1); x_6 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_6, 0, x_5); return x_6; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__17___lambda__3___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__17___lambda__3___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__17___lambda__2), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__17___lambda__2), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__17___lambda__3___closed__2() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__17___lambda__3___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__17___lambda__3___closed__1; +x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__17___lambda__3___closed__1; x_2 = lean_alloc_closure((void*)(l_Except_map___rarg), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__17___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__17___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; -x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__18(x_2); -x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__19(x_5, x_3, x_4); +x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__18(x_2); +x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__19(x_5, x_3, x_4); lean_dec(x_5); if (lean_obj_tag(x_6) == 0) { @@ -24785,7 +24785,7 @@ 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 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__17___lambda__3___closed__2; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__17___lambda__3___closed__2; x_13 = l_Task_Priority_default; x_14 = 0; x_15 = lean_task_map(x_12, x_11, x_13, x_14); @@ -24800,7 +24800,7 @@ x_17 = lean_ctor_get(x_9, 1); lean_inc(x_17); lean_inc(x_16); lean_dec(x_9); -x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__17___lambda__3___closed__2; +x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__17___lambda__3___closed__2; x_19 = l_Task_Priority_default; x_20 = 0; x_21 = lean_task_map(x_18, x_16, x_19, x_20); @@ -24859,28 +24859,28 @@ return x_30; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__17___lambda__4___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__17___lambda__4___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__17___lambda__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__17___lambda__1), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__17___lambda__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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__17___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_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_6 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__17___lambda__3), 4, 1); +x_6 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__17___lambda__3), 4, 1); lean_closure_set(x_6, 0, x_1); -x_7 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__4___closed__1; +x_7 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__4___closed__1; x_8 = lean_st_ref_take(x_7, x_5); 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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__17___lambda__4___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__17___lambda__4___closed__1; x_12 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_6); @@ -24907,12 +24907,12 @@ return x_18; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__17___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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__17___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; lean_object* x_7; uint8_t x_8; lean_dec(x_4); -x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__4___closed__1; +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__4___closed__1; x_7 = lean_st_ref_get(x_6, x_5); x_8 = !lean_is_exclusive(x_7); if (x_8 == 0) @@ -24927,7 +24927,7 @@ if (x_11 == 0) lean_object* x_12; lean_object* x_13; lean_free_object(x_7); x_12 = lean_box(0); -x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__17___lambda__4(x_1, x_2, x_3, x_12, x_10); +x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__17___lambda__4(x_1, x_2, x_3, x_12, x_10); return x_13; } else @@ -24935,10 +24935,10 @@ else lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_dec(x_2); lean_dec(x_1); -x_14 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_14 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_15 = lean_string_append(x_14, x_3); lean_dec(x_3); -x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__2; +x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__2; x_17 = lean_string_append(x_15, x_16); x_18 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_18, 0, x_17); @@ -24961,7 +24961,7 @@ if (x_21 == 0) { lean_object* x_22; lean_object* x_23; x_22 = lean_box(0); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__17___lambda__4(x_1, x_2, x_3, x_22, x_20); +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__17___lambda__4(x_1, x_2, x_3, x_22, x_20); return x_23; } else @@ -24969,10 +24969,10 @@ 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_dec(x_2); lean_dec(x_1); -x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_25 = lean_string_append(x_24, x_3); lean_dec(x_3); -x_26 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__2; +x_26 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__2; x_27 = lean_string_append(x_25, x_26); x_28 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_28, 0, x_27); @@ -24984,7 +24984,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__17(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__17(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; uint8_t x_7; @@ -25004,10 +25004,10 @@ if (x_8 == 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; x_9 = lean_ctor_get(x_5, 0); lean_dec(x_9); -x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_11 = lean_string_append(x_10, x_1); lean_dec(x_1); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___closed__1; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___closed__1; x_13 = lean_string_append(x_11, x_12); x_14 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_14, 0, x_13); @@ -25021,10 +25021,10 @@ lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean x_15 = lean_ctor_get(x_5, 1); lean_inc(x_15); lean_dec(x_5); -x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_17 = lean_string_append(x_16, x_1); lean_dec(x_1); -x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___closed__1; +x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___closed__1; x_19 = lean_string_append(x_17, x_18); x_20 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_20, 0, x_19); @@ -25041,12 +25041,12 @@ x_22 = lean_ctor_get(x_5, 1); lean_inc(x_22); lean_dec(x_5); x_23 = lean_box(0); -x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__17___lambda__5(x_2, x_3, x_1, x_23, x_22); +x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__17___lambda__5(x_2, x_3, x_1, x_23, x_22); return x_24; } } } -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__22(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__22(lean_object* x_1) { _start: { lean_object* x_2; @@ -25060,10 +25060,10 @@ if (x_3 == 0) lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; x_4 = lean_ctor_get(x_2, 0); x_5 = l_Lean_Json_compress(x_1); -x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__1; +x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__1; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); -x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__2; +x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__2; x_9 = lean_string_append(x_7, x_8); x_10 = lean_string_append(x_9, x_4); lean_dec(x_4); @@ -25083,10 +25083,10 @@ x_15 = lean_ctor_get(x_2, 0); lean_inc(x_15); lean_dec(x_2); x_16 = l_Lean_Json_compress(x_1); -x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__1; +x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__1; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); -x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__2; +x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__2; x_20 = lean_string_append(x_18, x_19); x_21 = lean_string_append(x_20, x_15); lean_dec(x_15); @@ -25123,7 +25123,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__23(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__23(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -25148,11 +25148,11 @@ return x_7; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__21___lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__21___lambda__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__22(x_1); +x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__22(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -25193,7 +25193,7 @@ return x_8; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__21___lambda__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__21___lambda__2(lean_object* x_1) { _start: { lean_object* x_2; size_t x_3; size_t x_4; lean_object* x_5; lean_object* x_6; @@ -25207,30 +25207,30 @@ lean_ctor_set(x_6, 0, x_5); return x_6; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__21___lambda__3___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__21___lambda__3___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__21___lambda__2), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__21___lambda__2), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__21___lambda__3___closed__2() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__21___lambda__3___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__21___lambda__3___closed__1; +x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__21___lambda__3___closed__1; x_2 = lean_alloc_closure((void*)(l_Except_map___rarg), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__21___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__21___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; -x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__22(x_2); -x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__23(x_5, x_3, x_4); +x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__22(x_2); +x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__23(x_5, x_3, x_4); lean_dec(x_5); if (lean_obj_tag(x_6) == 0) { @@ -25249,7 +25249,7 @@ 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 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__21___lambda__3___closed__2; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__21___lambda__3___closed__2; x_13 = l_Task_Priority_default; x_14 = 0; x_15 = lean_task_map(x_12, x_11, x_13, x_14); @@ -25264,7 +25264,7 @@ x_17 = lean_ctor_get(x_9, 1); lean_inc(x_17); lean_inc(x_16); lean_dec(x_9); -x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__21___lambda__3___closed__2; +x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__21___lambda__3___closed__2; x_19 = l_Task_Priority_default; x_20 = 0; x_21 = lean_task_map(x_18, x_16, x_19, x_20); @@ -25323,28 +25323,28 @@ return x_30; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__21___lambda__4___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__21___lambda__4___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__21___lambda__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__21___lambda__1), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____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_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__21___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_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_6 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__21___lambda__3), 4, 1); +x_6 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__21___lambda__3), 4, 1); lean_closure_set(x_6, 0, x_1); -x_7 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__4___closed__1; +x_7 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__4___closed__1; x_8 = lean_st_ref_take(x_7, x_5); 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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__21___lambda__4___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__21___lambda__4___closed__1; x_12 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_6); @@ -25371,12 +25371,12 @@ return x_18; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__21___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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__21___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; lean_object* x_7; uint8_t x_8; lean_dec(x_4); -x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__4___closed__1; +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__4___closed__1; x_7 = lean_st_ref_get(x_6, x_5); x_8 = !lean_is_exclusive(x_7); if (x_8 == 0) @@ -25391,7 +25391,7 @@ if (x_11 == 0) lean_object* x_12; lean_object* x_13; lean_free_object(x_7); x_12 = lean_box(0); -x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__21___lambda__4(x_1, x_2, x_3, x_12, x_10); +x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__21___lambda__4(x_1, x_2, x_3, x_12, x_10); return x_13; } else @@ -25399,10 +25399,10 @@ else lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_dec(x_2); lean_dec(x_1); -x_14 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_14 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_15 = lean_string_append(x_14, x_3); lean_dec(x_3); -x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__2; +x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__2; x_17 = lean_string_append(x_15, x_16); x_18 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_18, 0, x_17); @@ -25425,7 +25425,7 @@ if (x_21 == 0) { lean_object* x_22; lean_object* x_23; x_22 = lean_box(0); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__21___lambda__4(x_1, x_2, x_3, x_22, x_20); +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__21___lambda__4(x_1, x_2, x_3, x_22, x_20); return x_23; } else @@ -25433,10 +25433,10 @@ 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_dec(x_2); lean_dec(x_1); -x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_25 = lean_string_append(x_24, x_3); lean_dec(x_3); -x_26 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__2; +x_26 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__2; x_27 = lean_string_append(x_25, x_26); x_28 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_28, 0, x_27); @@ -25448,7 +25448,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__21(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__21(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; @@ -25468,10 +25468,10 @@ if (x_8 == 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; x_9 = lean_ctor_get(x_5, 0); lean_dec(x_9); -x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_11 = lean_string_append(x_10, x_1); lean_dec(x_1); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___closed__1; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___closed__1; x_13 = lean_string_append(x_11, x_12); x_14 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_14, 0, x_13); @@ -25485,10 +25485,10 @@ lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean x_15 = lean_ctor_get(x_5, 1); lean_inc(x_15); lean_dec(x_5); -x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_17 = lean_string_append(x_16, x_1); lean_dec(x_1); -x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___closed__1; +x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___closed__1; x_19 = lean_string_append(x_17, x_18); x_20 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_20, 0, x_19); @@ -25505,12 +25505,12 @@ x_22 = lean_ctor_get(x_5, 1); lean_inc(x_22); lean_dec(x_5); x_23 = lean_box(0); -x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__21___lambda__5(x_2, x_3, x_1, x_23, x_22); +x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__21___lambda__5(x_2, x_3, x_1, x_23, x_22); return x_24; } } } -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__25(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__25(lean_object* x_1) { _start: { lean_object* x_2; @@ -25524,10 +25524,10 @@ if (x_3 == 0) lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; x_4 = lean_ctor_get(x_2, 0); x_5 = l_Lean_Json_compress(x_1); -x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__1; +x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__1; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); -x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__2; +x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__2; x_9 = lean_string_append(x_7, x_8); x_10 = lean_string_append(x_9, x_4); lean_dec(x_4); @@ -25547,10 +25547,10 @@ x_15 = lean_ctor_get(x_2, 0); lean_inc(x_15); lean_dec(x_2); x_16 = l_Lean_Json_compress(x_1); -x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__1; +x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__1; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); -x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__2; +x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__2; x_20 = lean_string_append(x_18, x_19); x_21 = lean_string_append(x_20, x_15); lean_dec(x_15); @@ -25587,7 +25587,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__26(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__26(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -25612,11 +25612,11 @@ return x_7; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__24___lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__24___lambda__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__25(x_1); +x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__25(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -25657,7 +25657,7 @@ return x_8; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__24___lambda__2___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__24___lambda__2___closed__1() { _start: { lean_object* x_1; @@ -25665,22 +25665,22 @@ x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__L return x_1; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__24___lambda__2___closed__2() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__24___lambda__2___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__24___lambda__2___closed__1; +x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__24___lambda__2___closed__1; x_2 = lean_alloc_closure((void*)(l_Except_map___rarg), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__24___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__24___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; -x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__25(x_2); -x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__26(x_5, x_3, x_4); +x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__25(x_2); +x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__26(x_5, x_3, x_4); lean_dec(x_5); if (lean_obj_tag(x_6) == 0) { @@ -25699,7 +25699,7 @@ 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 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__24___lambda__2___closed__2; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__24___lambda__2___closed__2; x_13 = l_Task_Priority_default; x_14 = 0; x_15 = lean_task_map(x_12, x_11, x_13, x_14); @@ -25714,7 +25714,7 @@ x_17 = lean_ctor_get(x_9, 1); lean_inc(x_17); lean_inc(x_16); lean_dec(x_9); -x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__24___lambda__2___closed__2; +x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__24___lambda__2___closed__2; x_19 = l_Task_Priority_default; x_20 = 0; x_21 = lean_task_map(x_18, x_16, x_19, x_20); @@ -25773,28 +25773,28 @@ return x_30; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__24___lambda__3___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__24___lambda__3___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__24___lambda__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__24___lambda__1), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__24___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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__24___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_6 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__24___lambda__2), 4, 1); +x_6 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__24___lambda__2), 4, 1); lean_closure_set(x_6, 0, x_1); -x_7 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__4___closed__1; +x_7 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__4___closed__1; x_8 = lean_st_ref_take(x_7, x_5); 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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__24___lambda__3___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__24___lambda__3___closed__1; x_12 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_6); @@ -25821,12 +25821,12 @@ return x_18; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__24___lambda__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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__24___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_6; lean_object* x_7; uint8_t x_8; lean_dec(x_4); -x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__4___closed__1; +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__4___closed__1; x_7 = lean_st_ref_get(x_6, x_5); x_8 = !lean_is_exclusive(x_7); if (x_8 == 0) @@ -25841,7 +25841,7 @@ if (x_11 == 0) lean_object* x_12; lean_object* x_13; lean_free_object(x_7); x_12 = lean_box(0); -x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__24___lambda__3(x_1, x_2, x_3, x_12, x_10); +x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__24___lambda__3(x_1, x_2, x_3, x_12, x_10); return x_13; } else @@ -25849,10 +25849,10 @@ else lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_dec(x_2); lean_dec(x_1); -x_14 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_14 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_15 = lean_string_append(x_14, x_3); lean_dec(x_3); -x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__2; +x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__2; x_17 = lean_string_append(x_15, x_16); x_18 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_18, 0, x_17); @@ -25875,7 +25875,7 @@ if (x_21 == 0) { lean_object* x_22; lean_object* x_23; x_22 = lean_box(0); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__24___lambda__3(x_1, x_2, x_3, x_22, x_20); +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__24___lambda__3(x_1, x_2, x_3, x_22, x_20); return x_23; } else @@ -25883,10 +25883,10 @@ 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_dec(x_2); lean_dec(x_1); -x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_25 = lean_string_append(x_24, x_3); lean_dec(x_3); -x_26 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__2; +x_26 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__2; x_27 = lean_string_append(x_25, x_26); x_28 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_28, 0, x_27); @@ -25898,7 +25898,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__24(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__24(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; @@ -25918,10 +25918,10 @@ if (x_8 == 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; x_9 = lean_ctor_get(x_5, 0); lean_dec(x_9); -x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_11 = lean_string_append(x_10, x_1); lean_dec(x_1); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___closed__1; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___closed__1; x_13 = lean_string_append(x_11, x_12); x_14 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_14, 0, x_13); @@ -25935,10 +25935,10 @@ lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean x_15 = lean_ctor_get(x_5, 1); lean_inc(x_15); lean_dec(x_5); -x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_17 = lean_string_append(x_16, x_1); lean_dec(x_1); -x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___closed__1; +x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___closed__1; x_19 = lean_string_append(x_17, x_18); x_20 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_20, 0, x_19); @@ -25955,12 +25955,12 @@ x_22 = lean_ctor_get(x_5, 1); lean_inc(x_22); lean_dec(x_5); x_23 = lean_box(0); -x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__24___lambda__4(x_2, x_3, x_1, x_23, x_22); +x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__24___lambda__4(x_2, x_3, x_1, x_23, x_22); return x_24; } } } -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__28(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__28(lean_object* x_1) { _start: { lean_object* x_2; @@ -25974,10 +25974,10 @@ if (x_3 == 0) lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; x_4 = lean_ctor_get(x_2, 0); x_5 = l_Lean_Json_compress(x_1); -x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__1; +x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__1; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); -x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__2; +x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__2; x_9 = lean_string_append(x_7, x_8); x_10 = lean_string_append(x_9, x_4); lean_dec(x_4); @@ -25997,10 +25997,10 @@ x_15 = lean_ctor_get(x_2, 0); lean_inc(x_15); lean_dec(x_2); x_16 = l_Lean_Json_compress(x_1); -x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__1; +x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__1; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); -x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__2; +x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__2; x_20 = lean_string_append(x_18, x_19); x_21 = lean_string_append(x_20, x_15); lean_dec(x_15); @@ -26037,7 +26037,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__29(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__29(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -26062,11 +26062,11 @@ return x_7; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__27___lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__27___lambda__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__28(x_1); +x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__28(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -26116,12 +26116,12 @@ return x_11; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__27___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__27___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; -x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__28(x_2); -x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__29(x_5, x_3, x_4); +x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__28(x_2); +x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__29(x_5, x_3, x_4); lean_dec(x_5); if (lean_obj_tag(x_6) == 0) { @@ -26140,7 +26140,7 @@ 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 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__24___lambda__2___closed__2; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__24___lambda__2___closed__2; x_13 = l_Task_Priority_default; x_14 = 0; x_15 = lean_task_map(x_12, x_11, x_13, x_14); @@ -26155,7 +26155,7 @@ x_17 = lean_ctor_get(x_9, 1); lean_inc(x_17); lean_inc(x_16); lean_dec(x_9); -x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__24___lambda__2___closed__2; +x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__24___lambda__2___closed__2; x_19 = l_Task_Priority_default; x_20 = 0; x_21 = lean_task_map(x_18, x_16, x_19, x_20); @@ -26214,28 +26214,28 @@ return x_30; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__27___lambda__3___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__27___lambda__3___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__27___lambda__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__27___lambda__1), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__27___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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__27___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_6 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__27___lambda__2), 4, 1); +x_6 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__27___lambda__2), 4, 1); lean_closure_set(x_6, 0, x_1); -x_7 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__4___closed__1; +x_7 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__4___closed__1; x_8 = lean_st_ref_take(x_7, x_5); 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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__27___lambda__3___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__27___lambda__3___closed__1; x_12 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_6); @@ -26262,12 +26262,12 @@ return x_18; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__27___lambda__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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__27___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_6; lean_object* x_7; uint8_t x_8; lean_dec(x_4); -x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__4___closed__1; +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__4___closed__1; x_7 = lean_st_ref_get(x_6, x_5); x_8 = !lean_is_exclusive(x_7); if (x_8 == 0) @@ -26282,7 +26282,7 @@ if (x_11 == 0) lean_object* x_12; lean_object* x_13; lean_free_object(x_7); x_12 = lean_box(0); -x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__27___lambda__3(x_1, x_2, x_3, x_12, x_10); +x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__27___lambda__3(x_1, x_2, x_3, x_12, x_10); return x_13; } else @@ -26290,10 +26290,10 @@ else lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_dec(x_2); lean_dec(x_1); -x_14 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_14 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_15 = lean_string_append(x_14, x_3); lean_dec(x_3); -x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__2; +x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__2; x_17 = lean_string_append(x_15, x_16); x_18 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_18, 0, x_17); @@ -26316,7 +26316,7 @@ if (x_21 == 0) { lean_object* x_22; lean_object* x_23; x_22 = lean_box(0); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__27___lambda__3(x_1, x_2, x_3, x_22, x_20); +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__27___lambda__3(x_1, x_2, x_3, x_22, x_20); return x_23; } else @@ -26324,10 +26324,10 @@ 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_dec(x_2); lean_dec(x_1); -x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_25 = lean_string_append(x_24, x_3); lean_dec(x_3); -x_26 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__2; +x_26 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__2; x_27 = lean_string_append(x_25, x_26); x_28 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_28, 0, x_27); @@ -26339,7 +26339,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__27(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__27(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; @@ -26359,10 +26359,10 @@ if (x_8 == 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; x_9 = lean_ctor_get(x_5, 0); lean_dec(x_9); -x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_11 = lean_string_append(x_10, x_1); lean_dec(x_1); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___closed__1; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___closed__1; x_13 = lean_string_append(x_11, x_12); x_14 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_14, 0, x_13); @@ -26376,10 +26376,10 @@ lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean x_15 = lean_ctor_get(x_5, 1); lean_inc(x_15); lean_dec(x_5); -x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_17 = lean_string_append(x_16, x_1); lean_dec(x_1); -x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___closed__1; +x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___closed__1; x_19 = lean_string_append(x_17, x_18); x_20 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_20, 0, x_19); @@ -26396,12 +26396,12 @@ x_22 = lean_ctor_get(x_5, 1); lean_inc(x_22); lean_dec(x_5); x_23 = lean_box(0); -x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__27___lambda__4(x_2, x_3, x_1, x_23, x_22); +x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__27___lambda__4(x_2, x_3, x_1, x_23, x_22); return x_24; } } } -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__31(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__31(lean_object* x_1) { _start: { lean_object* x_2; @@ -26415,10 +26415,10 @@ if (x_3 == 0) lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; x_4 = lean_ctor_get(x_2, 0); x_5 = l_Lean_Json_compress(x_1); -x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__1; +x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__1; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); -x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__2; +x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__2; x_9 = lean_string_append(x_7, x_8); x_10 = lean_string_append(x_9, x_4); lean_dec(x_4); @@ -26438,10 +26438,10 @@ x_15 = lean_ctor_get(x_2, 0); lean_inc(x_15); lean_dec(x_2); x_16 = l_Lean_Json_compress(x_1); -x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__1; +x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__1; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); -x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__2; +x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__2; x_20 = lean_string_append(x_18, x_19); x_21 = lean_string_append(x_20, x_15); lean_dec(x_15); @@ -26478,7 +26478,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__32(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__32(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -26503,7 +26503,7 @@ return x_7; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__33(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__33(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -26528,11 +26528,11 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__30___lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__30___lambda__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__31(x_1); +x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__31(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -26573,7 +26573,7 @@ return x_8; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__30___lambda__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__30___lambda__2(lean_object* x_1) { _start: { lean_object* x_2; size_t x_3; size_t x_4; lean_object* x_5; lean_object* x_6; @@ -26581,36 +26581,36 @@ x_2 = lean_array_get_size(x_1); x_3 = lean_usize_of_nat(x_2); lean_dec(x_2); x_4 = 0; -x_5 = l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__33(x_3, x_4, x_1); +x_5 = l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__33(x_3, x_4, x_1); x_6 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_6, 0, x_5); return x_6; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__30___lambda__3___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__30___lambda__3___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__30___lambda__2), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__30___lambda__2), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__30___lambda__3___closed__2() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__30___lambda__3___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__30___lambda__3___closed__1; +x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__30___lambda__3___closed__1; x_2 = lean_alloc_closure((void*)(l_Except_map___rarg), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__30___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__30___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; -x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__31(x_2); -x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__32(x_5, x_3, x_4); +x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__31(x_2); +x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__32(x_5, x_3, x_4); lean_dec(x_5); if (lean_obj_tag(x_6) == 0) { @@ -26629,7 +26629,7 @@ 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 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__30___lambda__3___closed__2; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__30___lambda__3___closed__2; x_13 = l_Task_Priority_default; x_14 = 0; x_15 = lean_task_map(x_12, x_11, x_13, x_14); @@ -26644,7 +26644,7 @@ x_17 = lean_ctor_get(x_9, 1); lean_inc(x_17); lean_inc(x_16); lean_dec(x_9); -x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__30___lambda__3___closed__2; +x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__30___lambda__3___closed__2; x_19 = l_Task_Priority_default; x_20 = 0; x_21 = lean_task_map(x_18, x_16, x_19, x_20); @@ -26703,28 +26703,28 @@ return x_30; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__30___lambda__4___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__30___lambda__4___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__30___lambda__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__30___lambda__1), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__30___lambda__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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__30___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_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_6 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__30___lambda__3), 4, 1); +x_6 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__30___lambda__3), 4, 1); lean_closure_set(x_6, 0, x_1); -x_7 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__4___closed__1; +x_7 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__4___closed__1; x_8 = lean_st_ref_take(x_7, x_5); 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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__30___lambda__4___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__30___lambda__4___closed__1; x_12 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_6); @@ -26751,12 +26751,12 @@ return x_18; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__30___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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__30___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; lean_object* x_7; uint8_t x_8; lean_dec(x_4); -x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__4___closed__1; +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__4___closed__1; x_7 = lean_st_ref_get(x_6, x_5); x_8 = !lean_is_exclusive(x_7); if (x_8 == 0) @@ -26771,7 +26771,7 @@ if (x_11 == 0) lean_object* x_12; lean_object* x_13; lean_free_object(x_7); x_12 = lean_box(0); -x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__30___lambda__4(x_1, x_2, x_3, x_12, x_10); +x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__30___lambda__4(x_1, x_2, x_3, x_12, x_10); return x_13; } else @@ -26779,10 +26779,10 @@ else lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_dec(x_2); lean_dec(x_1); -x_14 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_14 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_15 = lean_string_append(x_14, x_3); lean_dec(x_3); -x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__2; +x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__2; x_17 = lean_string_append(x_15, x_16); x_18 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_18, 0, x_17); @@ -26805,7 +26805,7 @@ if (x_21 == 0) { lean_object* x_22; lean_object* x_23; x_22 = lean_box(0); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__30___lambda__4(x_1, x_2, x_3, x_22, x_20); +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__30___lambda__4(x_1, x_2, x_3, x_22, x_20); return x_23; } else @@ -26813,10 +26813,10 @@ 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_dec(x_2); lean_dec(x_1); -x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_25 = lean_string_append(x_24, x_3); lean_dec(x_3); -x_26 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__2; +x_26 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__2; x_27 = lean_string_append(x_25, x_26); x_28 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_28, 0, x_27); @@ -26828,7 +26828,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__30(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__30(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; @@ -26848,10 +26848,10 @@ if (x_8 == 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; x_9 = lean_ctor_get(x_5, 0); lean_dec(x_9); -x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_11 = lean_string_append(x_10, x_1); lean_dec(x_1); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___closed__1; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___closed__1; x_13 = lean_string_append(x_11, x_12); x_14 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_14, 0, x_13); @@ -26865,10 +26865,10 @@ lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean x_15 = lean_ctor_get(x_5, 1); lean_inc(x_15); lean_dec(x_5); -x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_17 = lean_string_append(x_16, x_1); lean_dec(x_1); -x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___closed__1; +x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___closed__1; x_19 = lean_string_append(x_17, x_18); x_20 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_20, 0, x_19); @@ -26885,12 +26885,12 @@ x_22 = lean_ctor_get(x_5, 1); lean_inc(x_22); lean_dec(x_5); x_23 = lean_box(0); -x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__30___lambda__5(x_2, x_3, x_1, x_23, x_22); +x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__30___lambda__5(x_2, x_3, x_1, x_23, x_22); return x_24; } } } -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__35(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__35(lean_object* x_1) { _start: { lean_object* x_2; @@ -26904,10 +26904,10 @@ if (x_3 == 0) lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; x_4 = lean_ctor_get(x_2, 0); x_5 = l_Lean_Json_compress(x_1); -x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__1; +x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__1; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); -x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__2; +x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__2; x_9 = lean_string_append(x_7, x_8); x_10 = lean_string_append(x_9, x_4); lean_dec(x_4); @@ -26927,10 +26927,10 @@ x_15 = lean_ctor_get(x_2, 0); lean_inc(x_15); lean_dec(x_2); x_16 = l_Lean_Json_compress(x_1); -x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__1; +x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__1; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); -x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__2; +x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__2; x_20 = lean_string_append(x_18, x_19); x_21 = lean_string_append(x_20, x_15); lean_dec(x_15); @@ -26967,7 +26967,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__36(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__36(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -26992,11 +26992,11 @@ return x_7; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__34___lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__34___lambda__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__35(x_1); +x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__35(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -27046,7 +27046,7 @@ return x_11; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__34___lambda__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__34___lambda__2(lean_object* x_1) { _start: { if (lean_obj_tag(x_1) == 0) @@ -27066,30 +27066,30 @@ return x_4; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__34___lambda__3___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__34___lambda__3___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__34___lambda__2), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__34___lambda__2), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__34___lambda__3___closed__2() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__34___lambda__3___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__34___lambda__3___closed__1; +x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__34___lambda__3___closed__1; x_2 = lean_alloc_closure((void*)(l_Except_map___rarg), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__34___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__34___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; -x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__35(x_2); -x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__36(x_5, x_3, x_4); +x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__35(x_2); +x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__36(x_5, x_3, x_4); lean_dec(x_5); if (lean_obj_tag(x_6) == 0) { @@ -27108,7 +27108,7 @@ 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 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__34___lambda__3___closed__2; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__34___lambda__3___closed__2; x_13 = l_Task_Priority_default; x_14 = 0; x_15 = lean_task_map(x_12, x_11, x_13, x_14); @@ -27123,7 +27123,7 @@ x_17 = lean_ctor_get(x_9, 1); lean_inc(x_17); lean_inc(x_16); lean_dec(x_9); -x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__34___lambda__3___closed__2; +x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__34___lambda__3___closed__2; x_19 = l_Task_Priority_default; x_20 = 0; x_21 = lean_task_map(x_18, x_16, x_19, x_20); @@ -27182,28 +27182,28 @@ return x_30; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__34___lambda__4___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__34___lambda__4___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__34___lambda__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__34___lambda__1), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__34___lambda__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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__34___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_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_6 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__34___lambda__3), 4, 1); +x_6 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__34___lambda__3), 4, 1); lean_closure_set(x_6, 0, x_1); -x_7 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__4___closed__1; +x_7 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__4___closed__1; x_8 = lean_st_ref_take(x_7, x_5); 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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__34___lambda__4___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__34___lambda__4___closed__1; x_12 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_6); @@ -27230,12 +27230,12 @@ return x_18; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__34___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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__34___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; lean_object* x_7; uint8_t x_8; lean_dec(x_4); -x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__4___closed__1; +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__4___closed__1; x_7 = lean_st_ref_get(x_6, x_5); x_8 = !lean_is_exclusive(x_7); if (x_8 == 0) @@ -27250,7 +27250,7 @@ if (x_11 == 0) lean_object* x_12; lean_object* x_13; lean_free_object(x_7); x_12 = lean_box(0); -x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__34___lambda__4(x_1, x_2, x_3, x_12, x_10); +x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__34___lambda__4(x_1, x_2, x_3, x_12, x_10); return x_13; } else @@ -27258,10 +27258,10 @@ else lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_dec(x_2); lean_dec(x_1); -x_14 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_14 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_15 = lean_string_append(x_14, x_3); lean_dec(x_3); -x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__2; +x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__2; x_17 = lean_string_append(x_15, x_16); x_18 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_18, 0, x_17); @@ -27284,7 +27284,7 @@ if (x_21 == 0) { lean_object* x_22; lean_object* x_23; x_22 = lean_box(0); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__34___lambda__4(x_1, x_2, x_3, x_22, x_20); +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__34___lambda__4(x_1, x_2, x_3, x_22, x_20); return x_23; } else @@ -27292,10 +27292,10 @@ 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_dec(x_2); lean_dec(x_1); -x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_25 = lean_string_append(x_24, x_3); lean_dec(x_3); -x_26 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__2; +x_26 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__2; x_27 = lean_string_append(x_25, x_26); x_28 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_28, 0, x_27); @@ -27307,7 +27307,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__34(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__34(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; @@ -27327,10 +27327,10 @@ if (x_8 == 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; x_9 = lean_ctor_get(x_5, 0); lean_dec(x_9); -x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_11 = lean_string_append(x_10, x_1); lean_dec(x_1); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___closed__1; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___closed__1; x_13 = lean_string_append(x_11, x_12); x_14 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_14, 0, x_13); @@ -27344,10 +27344,10 @@ lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean x_15 = lean_ctor_get(x_5, 1); lean_inc(x_15); lean_dec(x_5); -x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_17 = lean_string_append(x_16, x_1); lean_dec(x_1); -x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___closed__1; +x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___closed__1; x_19 = lean_string_append(x_17, x_18); x_20 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_20, 0, x_19); @@ -27364,12 +27364,12 @@ x_22 = lean_ctor_get(x_5, 1); lean_inc(x_22); lean_dec(x_5); x_23 = lean_box(0); -x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__34___lambda__5(x_2, x_3, x_1, x_23, x_22); +x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__34___lambda__5(x_2, x_3, x_1, x_23, x_22); return x_24; } } } -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__38(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__38(lean_object* x_1) { _start: { lean_object* x_2; @@ -27383,10 +27383,10 @@ if (x_3 == 0) lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; x_4 = lean_ctor_get(x_2, 0); x_5 = l_Lean_Json_compress(x_1); -x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__1; +x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__1; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); -x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__2; +x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__2; x_9 = lean_string_append(x_7, x_8); x_10 = lean_string_append(x_9, x_4); lean_dec(x_4); @@ -27406,10 +27406,10 @@ x_15 = lean_ctor_get(x_2, 0); lean_inc(x_15); lean_dec(x_2); x_16 = l_Lean_Json_compress(x_1); -x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__1; +x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__1; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); -x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__2; +x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__2; x_20 = lean_string_append(x_18, x_19); x_21 = lean_string_append(x_20, x_15); lean_dec(x_15); @@ -27446,7 +27446,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__39(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__39(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -27471,11 +27471,11 @@ return x_7; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__37___lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__37___lambda__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__38(x_1); +x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__38(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -27525,7 +27525,7 @@ return x_11; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__37___lambda__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__37___lambda__2(lean_object* x_1) { _start: { if (lean_obj_tag(x_1) == 0) @@ -27545,30 +27545,30 @@ return x_4; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__37___lambda__3___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__37___lambda__3___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__37___lambda__2), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__37___lambda__2), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__37___lambda__3___closed__2() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__37___lambda__3___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__37___lambda__3___closed__1; +x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__37___lambda__3___closed__1; x_2 = lean_alloc_closure((void*)(l_Except_map___rarg), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__37___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__37___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; -x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__38(x_2); -x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__39(x_5, x_3, x_4); +x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__38(x_2); +x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__39(x_5, x_3, x_4); lean_dec(x_5); if (lean_obj_tag(x_6) == 0) { @@ -27587,7 +27587,7 @@ 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 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__37___lambda__3___closed__2; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__37___lambda__3___closed__2; x_13 = l_Task_Priority_default; x_14 = 0; x_15 = lean_task_map(x_12, x_11, x_13, x_14); @@ -27602,7 +27602,7 @@ x_17 = lean_ctor_get(x_9, 1); lean_inc(x_17); lean_inc(x_16); lean_dec(x_9); -x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__37___lambda__3___closed__2; +x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__37___lambda__3___closed__2; x_19 = l_Task_Priority_default; x_20 = 0; x_21 = lean_task_map(x_18, x_16, x_19, x_20); @@ -27661,28 +27661,28 @@ return x_30; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__37___lambda__4___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__37___lambda__4___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__37___lambda__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__37___lambda__1), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__37___lambda__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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__37___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_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_6 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__37___lambda__3), 4, 1); +x_6 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__37___lambda__3), 4, 1); lean_closure_set(x_6, 0, x_1); -x_7 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__4___closed__1; +x_7 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__4___closed__1; x_8 = lean_st_ref_take(x_7, x_5); 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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__37___lambda__4___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__37___lambda__4___closed__1; x_12 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_6); @@ -27709,12 +27709,12 @@ return x_18; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__37___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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__37___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; lean_object* x_7; uint8_t x_8; lean_dec(x_4); -x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__4___closed__1; +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__4___closed__1; x_7 = lean_st_ref_get(x_6, x_5); x_8 = !lean_is_exclusive(x_7); if (x_8 == 0) @@ -27729,7 +27729,7 @@ if (x_11 == 0) lean_object* x_12; lean_object* x_13; lean_free_object(x_7); x_12 = lean_box(0); -x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__37___lambda__4(x_1, x_2, x_3, x_12, x_10); +x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__37___lambda__4(x_1, x_2, x_3, x_12, x_10); return x_13; } else @@ -27737,10 +27737,10 @@ else lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_dec(x_2); lean_dec(x_1); -x_14 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_14 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_15 = lean_string_append(x_14, x_3); lean_dec(x_3); -x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__2; +x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__2; x_17 = lean_string_append(x_15, x_16); x_18 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_18, 0, x_17); @@ -27763,7 +27763,7 @@ if (x_21 == 0) { lean_object* x_22; lean_object* x_23; x_22 = lean_box(0); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__37___lambda__4(x_1, x_2, x_3, x_22, x_20); +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__37___lambda__4(x_1, x_2, x_3, x_22, x_20); return x_23; } else @@ -27771,10 +27771,10 @@ 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_dec(x_2); lean_dec(x_1); -x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_25 = lean_string_append(x_24, x_3); lean_dec(x_3); -x_26 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__2; +x_26 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__2; x_27 = lean_string_append(x_25, x_26); x_28 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_28, 0, x_27); @@ -27786,7 +27786,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__37(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__37(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; @@ -27806,10 +27806,10 @@ if (x_8 == 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; x_9 = lean_ctor_get(x_5, 0); lean_dec(x_9); -x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_11 = lean_string_append(x_10, x_1); lean_dec(x_1); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___closed__1; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___closed__1; x_13 = lean_string_append(x_11, x_12); x_14 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_14, 0, x_13); @@ -27823,10 +27823,10 @@ lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean x_15 = lean_ctor_get(x_5, 1); lean_inc(x_15); lean_dec(x_5); -x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1; +x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1; x_17 = lean_string_append(x_16, x_1); lean_dec(x_1); -x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___closed__1; +x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___closed__1; x_19 = lean_string_append(x_17, x_18); x_20 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_20, 0, x_19); @@ -27843,12 +27843,12 @@ x_22 = lean_ctor_get(x_5, 1); lean_inc(x_22); lean_dec(x_5); x_23 = lean_box(0); -x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__37___lambda__5(x_2, x_3, x_1, x_23, x_22); +x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__37___lambda__5(x_2, x_3, x_1, x_23, x_22); return x_24; } } } -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____lambda__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____lambda__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; @@ -27859,7 +27859,7 @@ lean_ctor_set(x_4, 1, x_2); return x_4; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__1() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__1() { _start: { lean_object* x_1; @@ -27867,7 +27867,7 @@ x_1 = lean_mk_string_from_bytes("textDocument/waitForDiagnostics", 31); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__2() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__2() { _start: { lean_object* x_1; @@ -27875,15 +27875,15 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleWaitForDiagnosti return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__3() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__3() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____lambda__1___boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____lambda__1___boxed), 2, 0); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__4() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__4() { _start: { lean_object* x_1; @@ -27891,7 +27891,7 @@ x_1 = lean_mk_string_from_bytes("textDocument/completion", 23); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__5() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__5() { _start: { lean_object* x_1; @@ -27899,7 +27899,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleCompletion), 3, return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__6() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__6() { _start: { lean_object* x_1; @@ -27907,7 +27907,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Server_Completion_fillEligibleHeaderDecl return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__7() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__7() { _start: { lean_object* x_1; @@ -27915,7 +27915,7 @@ x_1 = lean_mk_string_from_bytes("completionItem/resolve", 22); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__8() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__8() { _start: { lean_object* x_1; @@ -27923,7 +27923,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleCompletionItemRe return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__9() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__9() { _start: { lean_object* x_1; @@ -27931,7 +27931,7 @@ x_1 = lean_mk_string_from_bytes("textDocument/hover", 18); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__10() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__10() { _start: { lean_object* x_1; @@ -27939,7 +27939,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleHover), 3, 0); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__11() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__11() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -27950,7 +27950,7 @@ lean_closure_set(x_3, 0, x_2); return x_3; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__12() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__12() { _start: { lean_object* x_1; @@ -27958,7 +27958,7 @@ x_1 = lean_mk_string_from_bytes("textDocument/declaration", 24); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__13() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__13() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -27969,7 +27969,7 @@ lean_closure_set(x_3, 0, x_2); return x_3; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__14() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__14() { _start: { lean_object* x_1; @@ -27977,7 +27977,7 @@ x_1 = lean_mk_string_from_bytes("textDocument/definition", 23); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__15() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__15() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -27988,7 +27988,7 @@ lean_closure_set(x_3, 0, x_2); return x_3; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__16() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__16() { _start: { lean_object* x_1; @@ -27996,7 +27996,7 @@ x_1 = lean_mk_string_from_bytes("textDocument/typeDefinition", 27); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__17() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__17() { _start: { lean_object* x_1; @@ -28004,7 +28004,7 @@ x_1 = lean_mk_string_from_bytes("textDocument/documentHighlight", 30); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__18() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__18() { _start: { lean_object* x_1; @@ -28012,7 +28012,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleDocumentHighligh return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__19() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__19() { _start: { lean_object* x_1; @@ -28020,7 +28020,7 @@ x_1 = lean_mk_string_from_bytes("textDocument/documentSymbol", 27); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__20() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__20() { _start: { lean_object* x_1; @@ -28028,7 +28028,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleDocumentSymbol__ return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__21() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__21() { _start: { lean_object* x_1; @@ -28036,7 +28036,7 @@ x_1 = lean_mk_string_from_bytes("textDocument/semanticTokens/full", 32); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__22() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__22() { _start: { lean_object* x_1; @@ -28044,7 +28044,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleSemanticTokensFu return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__23() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__23() { _start: { lean_object* x_1; @@ -28052,7 +28052,7 @@ x_1 = lean_mk_string_from_bytes("textDocument/semanticTokens/range", 33); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__24() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__24() { _start: { lean_object* x_1; @@ -28060,7 +28060,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleSemanticTokensRa return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__25() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__25() { _start: { lean_object* x_1; @@ -28068,7 +28068,7 @@ x_1 = lean_mk_string_from_bytes("textDocument/foldingRange", 25); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__26() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__26() { _start: { lean_object* x_1; @@ -28076,7 +28076,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleFoldingRange___b return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__27() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__27() { _start: { lean_object* x_1; @@ -28084,7 +28084,7 @@ x_1 = lean_mk_string_from_bytes("$/lean/plainGoal", 16); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__28() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__28() { _start: { lean_object* x_1; @@ -28092,7 +28092,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handlePlainGoal), 3, 0 return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__29() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__29() { _start: { lean_object* x_1; @@ -28100,7 +28100,7 @@ x_1 = lean_mk_string_from_bytes("$/lean/plainTermGoal", 20); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__30() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__30() { _start: { lean_object* x_1; @@ -28108,132 +28108,132 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handlePlainTermGoal), return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385_(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_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__1; -x_3 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__2; -x_4 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__3; -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1(x_2, x_3, x_4, x_1); +x_2 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__1; +x_3 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__2; +x_4 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__3; +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1(x_2, x_3, x_4, x_1); if (lean_obj_tag(x_5) == 0) { lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; x_6 = lean_ctor_get(x_5, 1); lean_inc(x_6); lean_dec(x_5); -x_7 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__4; -x_8 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__5; -x_9 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__6; -x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__4(x_7, x_8, x_9, x_6); +x_7 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__4; +x_8 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__5; +x_9 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__6; +x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__4(x_7, x_8, x_9, x_6); if (lean_obj_tag(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_10, 1); lean_inc(x_11); lean_dec(x_10); -x_12 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__7; -x_13 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__8; -x_14 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__7(x_12, x_13, x_4, x_11); +x_12 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__7; +x_13 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__8; +x_14 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__7(x_12, x_13, x_4, x_11); 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); -x_16 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__9; -x_17 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__10; -x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__10(x_16, x_17, x_4, x_15); +x_16 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__9; +x_17 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__10; +x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__10(x_16, x_17, x_4, x_15); 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, 1); lean_inc(x_19); lean_dec(x_18); -x_20 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__12; -x_21 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__11; -x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13(x_20, x_21, x_4, x_19); +x_20 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__12; +x_21 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__11; +x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13(x_20, x_21, x_4, x_19); if (lean_obj_tag(x_22) == 0) { lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; x_23 = lean_ctor_get(x_22, 1); lean_inc(x_23); lean_dec(x_22); -x_24 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__14; -x_25 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__13; -x_26 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13(x_24, x_25, x_4, x_23); +x_24 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__14; +x_25 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__13; +x_26 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13(x_24, x_25, x_4, x_23); if (lean_obj_tag(x_26) == 0) { lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; x_27 = lean_ctor_get(x_26, 1); lean_inc(x_27); lean_dec(x_26); -x_28 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__16; -x_29 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__15; -x_30 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13(x_28, x_29, x_4, x_27); +x_28 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__16; +x_29 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__15; +x_30 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13(x_28, x_29, x_4, x_27); if (lean_obj_tag(x_30) == 0) { lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; x_31 = lean_ctor_get(x_30, 1); lean_inc(x_31); lean_dec(x_30); -x_32 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__17; -x_33 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__18; -x_34 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__17(x_32, x_33, x_4, x_31); +x_32 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__17; +x_33 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__18; +x_34 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__17(x_32, x_33, x_4, x_31); if (lean_obj_tag(x_34) == 0) { lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; x_35 = lean_ctor_get(x_34, 1); lean_inc(x_35); lean_dec(x_34); -x_36 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__19; -x_37 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__20; -x_38 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__21(x_36, x_37, x_4, x_35); +x_36 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__19; +x_37 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__20; +x_38 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__21(x_36, x_37, x_4, x_35); if (lean_obj_tag(x_38) == 0) { lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; x_39 = lean_ctor_get(x_38, 1); lean_inc(x_39); lean_dec(x_38); -x_40 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__21; -x_41 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__22; -x_42 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__24(x_40, x_41, x_4, x_39); +x_40 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__21; +x_41 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__22; +x_42 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__24(x_40, x_41, x_4, x_39); if (lean_obj_tag(x_42) == 0) { 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); lean_dec(x_42); -x_44 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__23; -x_45 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__24; -x_46 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__27(x_44, x_45, x_4, x_43); +x_44 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__23; +x_45 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__24; +x_46 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__27(x_44, x_45, x_4, x_43); if (lean_obj_tag(x_46) == 0) { lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; x_47 = lean_ctor_get(x_46, 1); lean_inc(x_47); lean_dec(x_46); -x_48 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__25; -x_49 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__26; -x_50 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__30(x_48, x_49, x_4, x_47); +x_48 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__25; +x_49 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__26; +x_50 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__30(x_48, x_49, x_4, x_47); if (lean_obj_tag(x_50) == 0) { lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; x_51 = lean_ctor_get(x_50, 1); lean_inc(x_51); lean_dec(x_50); -x_52 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__27; -x_53 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__28; -x_54 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__34(x_52, x_53, x_4, x_51); +x_52 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__27; +x_53 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__28; +x_54 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__34(x_52, x_53, x_4, x_51); if (lean_obj_tag(x_54) == 0) { lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; x_55 = lean_ctor_get(x_54, 1); lean_inc(x_55); lean_dec(x_54); -x_56 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__29; -x_57 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__30; -x_58 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__37(x_56, x_57, x_4, x_55); +x_56 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__29; +x_57 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__30; +x_58 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__37(x_56, x_57, x_4, x_55); return x_58; } else @@ -28536,102 +28536,102 @@ return x_110; } } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__3(x_1, x_2, x_3); +x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__3(x_1, x_2, x_3); lean_dec(x_2); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__2___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__2___boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__2(x_1); +x_2 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__2(x_1); lean_dec(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____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_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____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) { _start: { lean_object* x_6; -x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__4(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__4(x_1, x_2, x_3, x_4, x_5); lean_dec(x_4); return x_6; } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__6(x_1, x_2, x_3); +x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__6(x_1, x_2, x_3); lean_dec(x_2); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____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_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____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) { _start: { lean_object* x_6; -x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__4___lambda__3(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__4___lambda__3(x_1, x_2, x_3, x_4, x_5); lean_dec(x_4); return x_6; } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__9(x_1, x_2, x_3); +x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__9(x_1, x_2, x_3); lean_dec(x_2); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____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_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____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) { _start: { lean_object* x_6; -x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__7___lambda__3(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__7___lambda__3(x_1, x_2, x_3, x_4, x_5); lean_dec(x_4); return x_6; } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__12(x_1, x_2, x_3); +x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__12(x_1, x_2, x_3); lean_dec(x_2); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__10___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__10___lambda__4___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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__10___lambda__4(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__10___lambda__4(x_1, x_2, x_3, x_4, x_5); lean_dec(x_4); return x_6; } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__15___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__15___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__15(x_1, x_2, x_3); +x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__15(x_1, x_2, x_3); lean_dec(x_2); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__16___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__16___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; @@ -28639,30 +28639,30 @@ x_4 = lean_unbox_usize(x_1); lean_dec(x_1); x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__16(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__16(x_4, x_5, x_3); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13___lambda__4___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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13___lambda__4(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13___lambda__4(x_1, x_2, x_3, x_4, x_5); lean_dec(x_4); return x_6; } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__19___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__19___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__19(x_1, x_2, x_3); +x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__19(x_1, x_2, x_3); lean_dec(x_2); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__20___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__20___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; @@ -28670,87 +28670,87 @@ x_4 = lean_unbox_usize(x_1); lean_dec(x_1); x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__20(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__20(x_4, x_5, x_3); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__17___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__17___lambda__4___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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__17___lambda__4(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__17___lambda__4(x_1, x_2, x_3, x_4, x_5); lean_dec(x_4); return x_6; } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__23___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__23___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__23(x_1, x_2, x_3); +x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__23(x_1, x_2, x_3); lean_dec(x_2); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__21___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__21___lambda__4___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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__21___lambda__4(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__21___lambda__4(x_1, x_2, x_3, x_4, x_5); lean_dec(x_4); return x_6; } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__26___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__26___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__26(x_1, x_2, x_3); +x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__26(x_1, x_2, x_3); lean_dec(x_2); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__24___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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__24___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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__24___lambda__3(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__24___lambda__3(x_1, x_2, x_3, x_4, x_5); lean_dec(x_4); return x_6; } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__29___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__29___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__29(x_1, x_2, x_3); +x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__29(x_1, x_2, x_3); lean_dec(x_2); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__27___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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__27___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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__27___lambda__3(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__27___lambda__3(x_1, x_2, x_3, x_4, x_5); lean_dec(x_4); return x_6; } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__32___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__32___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__32(x_1, x_2, x_3); +x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__32(x_1, x_2, x_3); lean_dec(x_2); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__33___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__33___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; @@ -28758,62 +28758,62 @@ x_4 = lean_unbox_usize(x_1); lean_dec(x_1); x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__33(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__33(x_4, x_5, x_3); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__30___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__30___lambda__4___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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__30___lambda__4(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__30___lambda__4(x_1, x_2, x_3, x_4, x_5); lean_dec(x_4); return x_6; } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__36___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__36___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__36(x_1, x_2, x_3); +x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__36(x_1, x_2, x_3); lean_dec(x_2); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__34___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__34___lambda__4___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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__34___lambda__4(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__34___lambda__4(x_1, x_2, x_3, x_4, x_5); lean_dec(x_4); return x_6; } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__39___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__39___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__39(x_1, x_2, x_3); +x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__39(x_1, x_2, x_3); lean_dec(x_2); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__37___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__37___lambda__4___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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__37___lambda__4(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__37___lambda__4(x_1, x_2, x_3, x_4, x_5); lean_dec(x_4); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____lambda__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____lambda__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____lambda__1(x_1, x_2); +x_3 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____lambda__1(x_1, x_2); lean_dec(x_1); return x_3; } @@ -29141,149 +29141,149 @@ l_Lean_Server_FileWorker_handleWaitForDiagnostics___lambda__2___closed__1 = _ini lean_mark_persistent(l_Lean_Server_FileWorker_handleWaitForDiagnostics___lambda__2___closed__1); l_Lean_Server_FileWorker_handleWaitForDiagnostics___closed__1 = _init_l_Lean_Server_FileWorker_handleWaitForDiagnostics___closed__1(); lean_mark_persistent(l_Lean_Server_FileWorker_handleWaitForDiagnostics___closed__1); -l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__1 = _init_l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__1(); -lean_mark_persistent(l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__1); -l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__2 = _init_l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__2(); -lean_mark_persistent(l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__2___closed__2); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__2___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__2___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__2___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__3___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__3___closed__2(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__3___closed__2); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__4___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__4___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__4___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__4___closed__2(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__4___closed__2); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__2(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___lambda__5___closed__2); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__1___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__4___lambda__2___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__4___lambda__2___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__4___lambda__2___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__4___lambda__2___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__4___lambda__2___closed__2(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__4___lambda__2___closed__2); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__4___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__4___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__4___lambda__3___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__7___lambda__2___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__7___lambda__2___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__7___lambda__2___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__7___lambda__2___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__7___lambda__2___closed__2(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__7___lambda__2___closed__2); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__7___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__7___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__7___lambda__3___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__10___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__10___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__10___lambda__3___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__10___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__10___lambda__3___closed__2(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__10___lambda__3___closed__2); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__10___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__10___lambda__4___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__10___lambda__4___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13___lambda__3___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13___lambda__3___closed__2(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13___lambda__3___closed__2); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13___lambda__4___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__13___lambda__4___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__17___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__17___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__17___lambda__3___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__17___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__17___lambda__3___closed__2(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__17___lambda__3___closed__2); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__17___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__17___lambda__4___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__17___lambda__4___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__21___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__21___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__21___lambda__3___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__21___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__21___lambda__3___closed__2(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__21___lambda__3___closed__2); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__21___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__21___lambda__4___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__21___lambda__4___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__24___lambda__2___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__24___lambda__2___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__24___lambda__2___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__24___lambda__2___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__24___lambda__2___closed__2(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__24___lambda__2___closed__2); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__24___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__24___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__24___lambda__3___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__27___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__27___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__27___lambda__3___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__30___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__30___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__30___lambda__3___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__30___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__30___lambda__3___closed__2(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__30___lambda__3___closed__2); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__30___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__30___lambda__4___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__30___lambda__4___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__34___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__34___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__34___lambda__3___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__34___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__34___lambda__3___closed__2(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__34___lambda__3___closed__2); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__34___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__34___lambda__4___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__34___lambda__4___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__37___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__37___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__37___lambda__3___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__37___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__37___lambda__3___closed__2(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__37___lambda__3___closed__2); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__37___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__37___lambda__4___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____spec__37___lambda__4___closed__1); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__1 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__1(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__1); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__2 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__2(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__2); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__3 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__3(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__3); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__4 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__4(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__4); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__5 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__5(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__5); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__6 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__6(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__6); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__7 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__7(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__7); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__8 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__8(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__8); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__9 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__9(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__9); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__10 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__10(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__10); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__11 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__11(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__11); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__12 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__12(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__12); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__13 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__13(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__13); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__14 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__14(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__14); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__15 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__15(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__15); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__16 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__16(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__16); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__17 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__17(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__17); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__18 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__18(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__18); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__19 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__19(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__19); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__20 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__20(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__20); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__21 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__21(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__21); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__22 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__22(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__22); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__23 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__23(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__23); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__24 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__24(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__24); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__25 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__25(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__25); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__26 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__26(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__26); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__27 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__27(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__27); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__28 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__28(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__28); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__29 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__29(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__29); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__30 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__30(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383____closed__30); -if (builtin) {res = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11383_(lean_io_mk_world()); +l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__1 = _init_l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__1(); +lean_mark_persistent(l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__1); +l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__2 = _init_l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__2(); +lean_mark_persistent(l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__2___closed__2); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__2___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__2___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__3___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__3___closed__2); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__4___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__4___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__4___closed__2(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__4___closed__2); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__2(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___lambda__5___closed__2); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__1___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__4___lambda__2___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__4___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__4___lambda__2___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__4___lambda__2___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__4___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__4___lambda__2___closed__2); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__4___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__4___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__4___lambda__3___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__7___lambda__2___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__7___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__7___lambda__2___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__7___lambda__2___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__7___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__7___lambda__2___closed__2); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__7___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__7___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__7___lambda__3___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__10___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__10___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__10___lambda__3___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__10___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__10___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__10___lambda__3___closed__2); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__10___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__10___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__10___lambda__4___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13___lambda__3___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13___lambda__3___closed__2); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__13___lambda__4___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__17___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__17___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__17___lambda__3___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__17___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__17___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__17___lambda__3___closed__2); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__17___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__17___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__17___lambda__4___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__21___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__21___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__21___lambda__3___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__21___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__21___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__21___lambda__3___closed__2); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__21___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__21___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__21___lambda__4___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__24___lambda__2___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__24___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__24___lambda__2___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__24___lambda__2___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__24___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__24___lambda__2___closed__2); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__24___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__24___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__24___lambda__3___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__27___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__27___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__27___lambda__3___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__30___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__30___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__30___lambda__3___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__30___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__30___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__30___lambda__3___closed__2); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__30___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__30___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__30___lambda__4___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__34___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__34___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__34___lambda__3___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__34___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__34___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__34___lambda__3___closed__2); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__34___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__34___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__34___lambda__4___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__37___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__37___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__37___lambda__3___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__37___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__37___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__37___lambda__3___closed__2); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__37___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__37___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____spec__37___lambda__4___closed__1); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__1 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__1(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__1); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__2 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__2(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__2); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__3 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__3(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__3); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__4 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__4(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__4); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__5 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__5(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__5); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__6 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__6(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__6); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__7 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__7(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__7); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__8 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__8(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__8); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__9 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__9(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__9); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__10 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__10(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__10); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__11 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__11(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__11); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__12 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__12(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__12); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__13 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__13(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__13); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__14 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__14(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__14); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__15 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__15(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__15); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__16 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__16(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__16); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__17 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__17(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__17); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__18 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__18(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__18); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__19 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__19(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__19); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__20 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__20(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__20); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__21 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__21(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__21); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__22 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__22(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__22); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__23 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__23(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__23); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__24 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__24(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__24); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__25 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__25(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__25); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__26 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__26(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__26); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__27 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__27(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__27); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__28 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__28(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__28); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__29 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__29(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__29); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__30 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__30(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385____closed__30); +if (builtin) {res = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_11385_(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/References.c b/stage0/stdlib/Lean/Server/References.c index 18e427de7c58..36c71325dd73 100644 --- a/stage0/stdlib/Lean/Server/References.c +++ b/stage0/stdlib/Lean/Server/References.c @@ -5546,7 +5546,7 @@ static lean_object* _init_l_Lean_Loop_forIn_loop___at_Lean_Server_combineIdents_ 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_Loop_forIn_loop___at_Lean_Server_combineIdents_findCanonicalRepresentative___spec__6___closed__1; x_2 = l_Lean_Loop_forIn_loop___at_Lean_Server_combineIdents_findCanonicalRepresentative___spec__6___closed__2; -x_3 = lean_unsigned_to_nat(184u); +x_3 = lean_unsigned_to_nat(210u); x_4 = lean_unsigned_to_nat(14u); x_5 = l_Lean_Loop_forIn_loop___at_Lean_Server_combineIdents_findCanonicalRepresentative___spec__6___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); From 9ee1ff243528a377243fb2434b2a410a9355eb3b Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Wed, 13 Mar 2024 20:46:04 -0700 Subject: [PATCH 30/32] chore: remove bootstrapping workaround --- src/Lean/Elab/Declaration.lean | 1 - src/Lean/Elab/DefView.lean | 3 --- src/Lean/Parser/Command.lean | 4 +--- 3 files changed, 1 insertion(+), 7 deletions(-) diff --git a/src/Lean/Elab/Declaration.lean b/src/Lean/Elab/Declaration.lean index 90e2cf682b12..70ea51f3548c 100644 --- a/src/Lean/Elab/Declaration.lean +++ b/src/Lean/Elab/Declaration.lean @@ -43,7 +43,6 @@ private def isNamedDef (stx : Syntax) : Bool := let k := decl.getKind k == ``Lean.Parser.Command.abbrev || k == ``Lean.Parser.Command.definition || - k == ``Lean.Parser.Command.def || -- TODO: delete k == ``Lean.Parser.Command.theorem || k == ``Lean.Parser.Command.opaque || k == ``Lean.Parser.Command.axiom || diff --git a/src/Lean/Elab/DefView.lean b/src/Lean/Elab/DefView.lean index 71c1f9242986..1ac40b2bfd3f 100644 --- a/src/Lean/Elab/DefView.lean +++ b/src/Lean/Elab/DefView.lean @@ -143,7 +143,6 @@ def isDefLike (stx : Syntax) : Bool := let declKind := stx.getKind declKind == ``Parser.Command.abbrev || declKind == ``Parser.Command.definition || - declKind == ``Parser.Command.def || -- TODO: delete declKind == ``Parser.Command.theorem || declKind == ``Parser.Command.opaque || declKind == ``Parser.Command.instance || @@ -155,8 +154,6 @@ def mkDefView (modifiers : Modifiers) (stx : Syntax) : CommandElabM DefView := return mkDefViewOfAbbrev modifiers stx else if declKind == ``Parser.Command.definition then return mkDefViewOfDef modifiers stx - else if declKind == ``Parser.Command.def then -- TODO: delete - return mkDefViewOfDef modifiers stx else if declKind == ``Parser.Command.theorem then return mkDefViewOfTheorem modifiers stx else if declKind == ``Parser.Command.opaque then diff --git a/src/Lean/Parser/Command.lean b/src/Lean/Parser/Command.lean index 995d55d43fb4..1035dc9aada5 100644 --- a/src/Lean/Parser/Command.lean +++ b/src/Lean/Parser/Command.lean @@ -116,8 +116,6 @@ def optDefDeriving := optional (ppDedent ppLine >> atomic ("deriving " >> notSymbol "instance") >> sepBy1 ident ", ") def definition := leading_parser "def " >> recover declId skipUntilWsOrDelim >> ppIndent optDeclSig >> declVal >> optDefDeriving -def «def» := leading_parser -- TODO: delete - "def " >> recover declId skipUntilWsOrDelim >> ppIndent optDeclSig >> declVal >> optDefDeriving def «theorem» := leading_parser "theorem " >> recover declId skipUntilWsOrDelim >> ppIndent declSig >> declVal def «opaque» := leading_parser @@ -200,7 +198,7 @@ def «structure» := leading_parser optDeriving @[builtin_command_parser] def declaration := leading_parser declModifiers false >> - («abbrev» <|> definition <|> «def» <|> «theorem» <|> «opaque» <|> «instance» <|> «axiom» <|> «example» <|> + («abbrev» <|> definition <|> «theorem» <|> «opaque» <|> «instance» <|> «axiom» <|> «example» <|> «inductive» <|> classInductive <|> «structure») @[builtin_command_parser] def «deriving» := leading_parser "deriving " >> "instance " >> derivingClasses >> " for " >> sepBy1 (recover ident skip) ", " From 214179b6b9315f47b9d36c6ebe044edab241b489 Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Wed, 13 Mar 2024 20:47:18 -0700 Subject: [PATCH 31/32] chore: update stage0 --- stage0/stdlib/Init/Meta.c | 881 +++--- stage0/stdlib/Init/NotationExtra.c | 187 +- stage0/stdlib/Init/Simproc.c | 187 +- stage0/stdlib/Lean/Elab/Declaration.c | 500 ++-- stage0/stdlib/Lean/Elab/DefView.c | 335 +-- stage0/stdlib/Lean/Elab/Quotation.c | 189 +- stage0/stdlib/Lean/Elab/Tactic/Config.c | 1023 +++---- .../Meta/Tactic/Simp/BuiltinSimprocs/UInt.c | 2571 +++++++++-------- stage0/stdlib/Lean/Parser/Command.c | 1183 +++----- 9 files changed, 3359 insertions(+), 3697 deletions(-) diff --git a/stage0/stdlib/Init/Meta.c b/stage0/stdlib/Init/Meta.c index 963d81d022cf..a1a2529e2e57 100644 --- a/stage0/stdlib/Init/Meta.c +++ b/stage0/stdlib/Init/Meta.c @@ -800,6 +800,7 @@ static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Me lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(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__51; 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___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__143; 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; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__123; @@ -23125,7 +23126,7 @@ static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("def", 3); +x_1 = lean_mk_string_from_bytes("definition", 10); return x_1; } } @@ -23145,30 +23146,38 @@ static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("expandSimp", 10); +x_1 = lean_mk_string_from_bytes("def", 3); return x_1; } } static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__18() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("expandSimp", 10); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__19() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__17; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__18; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__19() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__17; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__18; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__20() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__21() { _start: { lean_object* x_1; @@ -23176,19 +23185,19 @@ x_1 = lean_mk_string_from_bytes("optDeclSig", 10); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__21() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__22() { _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_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__2; x_3 = l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStr4Nil___closed__1; -x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__20; +x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__21; 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___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__22() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__23() { _start: { lean_object* x_1; @@ -23196,19 +23205,19 @@ x_1 = lean_mk_string_from_bytes("typeSpec", 8); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__23() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__24() { _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_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__2; x_3 = l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__3; -x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__22; +x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__23; 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___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__24() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__25() { _start: { lean_object* x_1; @@ -23216,82 +23225,82 @@ x_1 = lean_mk_string_from_bytes("Macro", 5); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__25() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__26() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__24; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__25; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__26() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__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_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__24; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__25; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__27() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__28() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__1; -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__24; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__25; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__28() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__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_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__27; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__28; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__29() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__30() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__27; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___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_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__30() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__31() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__29; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__30; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__31() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__32() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__28; -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__30; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__29; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__31; 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___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__32() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__33() { _start: { lean_object* x_1; @@ -23299,19 +23308,19 @@ x_1 = lean_mk_string_from_bytes("declValSimple", 13); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__33() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__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_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_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStr4Nil___closed__1; -x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__32; +x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___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_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__34() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__35() { _start: { lean_object* x_1; @@ -23319,19 +23328,19 @@ x_1 = lean_mk_string_from_bytes("fun", 3); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__35() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__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_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_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__3; -x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__34; +x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__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_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__36() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__37() { _start: { lean_object* x_1; @@ -23339,19 +23348,19 @@ x_1 = lean_mk_string_from_bytes("basicFun", 8); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__37() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___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_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_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__3; -x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__36; +x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___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_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__38() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__39() { _start: { lean_object* x_1; @@ -23359,26 +23368,26 @@ x_1 = lean_mk_string_from_bytes("s", 1); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__39() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__40() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__38; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__39; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__40() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__41() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__38; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__39; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__41() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__42() { _start: { lean_object* x_1; @@ -23386,7 +23395,7 @@ x_1 = lean_mk_string_from_bytes("=>", 2); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__42() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__43() { _start: { lean_object* x_1; @@ -23394,19 +23403,19 @@ x_1 = lean_mk_string_from_bytes("do", 2); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__43() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___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_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_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__3; -x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__42; +x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___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_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__44() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__45() { _start: { lean_object* x_1; @@ -23414,19 +23423,19 @@ x_1 = lean_mk_string_from_bytes("doSeqIndent", 11); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__45() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__46() { _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_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__2; x_3 = l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__3; -x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__44; +x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__45; 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___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__46() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__47() { _start: { lean_object* x_1; @@ -23434,19 +23443,19 @@ x_1 = lean_mk_string_from_bytes("doSeqItem", 9); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__47() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__48() { _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_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__2; x_3 = l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__3; -x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__46; +x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__47; 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___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__48() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__49() { _start: { lean_object* x_1; @@ -23454,19 +23463,19 @@ x_1 = lean_mk_string_from_bytes("doLetArrow", 10); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__49() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___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_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_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__3; -x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__48; +x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___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_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__50() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__51() { _start: { lean_object* x_1; @@ -23474,7 +23483,7 @@ x_1 = lean_mk_string_from_bytes("let", 3); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__51() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__52() { _start: { lean_object* x_1; @@ -23482,19 +23491,19 @@ x_1 = lean_mk_string_from_bytes("doIdDecl", 8); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__52() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___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_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_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__3; -x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__51; +x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___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_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__53() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__54() { _start: { lean_object* x_1; @@ -23502,26 +23511,26 @@ x_1 = lean_mk_string_from_bytes("c", 1); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__54() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__55() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__53; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__54; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__55() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__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_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__53; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__54; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__56() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__57() { _start: { lean_object* x_1; @@ -23529,7 +23538,7 @@ x_1 = lean_mk_string_from_bytes("←", 3); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__57() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__58() { _start: { lean_object* x_1; @@ -23537,19 +23546,19 @@ x_1 = lean_mk_string_from_bytes("doMatch", 7); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__58() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___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_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_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__3; -x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__57; +x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___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_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__59() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__60() { _start: { lean_object* x_1; @@ -23557,7 +23566,7 @@ x_1 = lean_mk_string_from_bytes("match", 5); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__60() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__61() { _start: { lean_object* x_1; @@ -23565,19 +23574,19 @@ x_1 = lean_mk_string_from_bytes("matchDiscr", 10); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__61() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__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_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_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__3; -x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__60; +x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__61; 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___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__62() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__63() { _start: { lean_object* x_1; @@ -23585,17 +23594,17 @@ x_1 = lean_mk_string_from_bytes("term__[_]", 9); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__63() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__64() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__62; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__63; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__64() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__65() { _start: { lean_object* x_1; @@ -23603,7 +23612,7 @@ x_1 = lean_mk_string_from_bytes("1", 1); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__65() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__66() { _start: { lean_object* x_1; @@ -23611,7 +23620,7 @@ x_1 = lean_mk_string_from_bytes("0", 1); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__66() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__67() { _start: { lean_object* x_1; @@ -23619,7 +23628,7 @@ x_1 = lean_mk_string_from_bytes("with", 4); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__67() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__68() { _start: { lean_object* x_1; @@ -23627,19 +23636,19 @@ x_1 = lean_mk_string_from_bytes("matchAlts", 9); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__68() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__69() { _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_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__2; x_3 = l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__3; -x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__67; +x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__68; 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___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__69() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__70() { _start: { lean_object* x_1; @@ -23647,19 +23656,19 @@ x_1 = lean_mk_string_from_bytes("matchAlt", 8); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__70() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__71() { _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_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__2; x_3 = l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__3; -x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__69; +x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__70; 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___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__71() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__72() { _start: { lean_object* x_1; @@ -23667,7 +23676,7 @@ x_1 = lean_mk_string_from_bytes("|", 1); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__72() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__73() { _start: { lean_object* x_1; @@ -23675,19 +23684,19 @@ x_1 = lean_mk_string_from_bytes("dynamicQuot", 11); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__73() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__74() { _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_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__2; x_3 = l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__3; -x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__72; +x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__73; 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___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__74() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__75() { _start: { lean_object* x_1; @@ -23695,7 +23704,7 @@ x_1 = lean_mk_string_from_bytes("`(", 2); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__75() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__76() { _start: { lean_object* x_1; lean_object* x_2; @@ -23704,7 +23713,7 @@ x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__76() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__77() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -23714,7 +23723,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__77() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__78() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -23726,7 +23735,7 @@ lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__78() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__79() { _start: { lean_object* x_1; lean_object* x_2; @@ -23736,31 +23745,31 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__79() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__80() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__78; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__79; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__80() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__81() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__77; -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__79; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__78; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__80; 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___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__81() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__82() { _start: { lean_object* x_1; @@ -23768,7 +23777,7 @@ x_1 = lean_mk_string_from_bytes("pseudo", 6); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__82() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__83() { _start: { lean_object* x_1; @@ -23776,18 +23785,18 @@ x_1 = lean_mk_string_from_bytes("antiquot", 8); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__83() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__84() { _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__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_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__82; +x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__83; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__84() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__85() { _start: { lean_object* x_1; @@ -23795,7 +23804,7 @@ x_1 = lean_mk_string_from_bytes("$", 1); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__85() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__86() { _start: { lean_object* x_1; @@ -23803,19 +23812,19 @@ x_1 = lean_mk_string_from_bytes("doExpr", 6); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__86() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__87() { _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_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__2; x_3 = l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__3; -x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__85; +x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__86; 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___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__87() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__88() { _start: { lean_object* x_1; @@ -23823,17 +23832,17 @@ x_1 = lean_mk_string_from_bytes("choice", 6); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__88() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__89() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__87; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__88; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__89() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__90() { _start: { lean_object* x_1; @@ -23841,17 +23850,17 @@ x_1 = lean_mk_string_from_bytes("term{}", 6); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__90() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__91() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__89; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__90; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__91() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__92() { _start: { lean_object* x_1; @@ -23859,19 +23868,19 @@ x_1 = lean_mk_string_from_bytes("doLet", 5); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__92() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__93() { _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_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__2; x_3 = l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__3; -x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__91; +x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__92; 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___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__93() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__94() { _start: { lean_object* x_1; @@ -23879,19 +23888,19 @@ x_1 = lean_mk_string_from_bytes("letDecl", 7); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__94() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__95() { _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_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__2; x_3 = l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__3; -x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__93; +x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__94; 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___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__95() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__96() { _start: { lean_object* x_1; @@ -23899,19 +23908,19 @@ x_1 = lean_mk_string_from_bytes("letIdDecl", 9); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__96() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__97() { _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_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__2; x_3 = l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__3; -x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__95; +x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__96; 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___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__97() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__98() { _start: { lean_object* x_1; @@ -23919,16 +23928,16 @@ x_1 = lean_mk_string_from_bytes("s.setKind", 9); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__98() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__99() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__97; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__98; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__99() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__100() { _start: { lean_object* x_1; @@ -23936,17 +23945,17 @@ x_1 = lean_mk_string_from_bytes("setKind", 7); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__100() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__101() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__38; -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__99; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__39; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__100; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__101() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__102() { _start: { lean_object* x_1; @@ -23954,16 +23963,16 @@ x_1 = lean_mk_string_from_bytes("s.setArg", 8); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__102() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__103() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__101; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__102; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__103() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__104() { _start: { lean_object* x_1; @@ -23971,17 +23980,17 @@ x_1 = lean_mk_string_from_bytes("setArg", 6); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__104() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__105() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__38; -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__103; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__39; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__104; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__105() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__106() { _start: { lean_object* x_1; @@ -23989,19 +23998,19 @@ x_1 = lean_mk_string_from_bytes("paren", 5); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__106() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__107() { _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_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__2; x_3 = l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__3; -x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__105; +x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__106; 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___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__107() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__108() { _start: { lean_object* x_1; @@ -24009,82 +24018,82 @@ x_1 = lean_mk_string_from_bytes("mkAtomFrom", 10); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__108() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__109() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__107; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__108; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__109() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__110() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__107; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__108; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__110() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__111() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__1; -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__107; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__108; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__111() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__112() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__110; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__111; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__112() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__113() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__110; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___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_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__113() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__114() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__112; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__113; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__114() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__115() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__111; -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__113; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__112; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__114; 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___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__115() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__116() { _start: { lean_object* x_1; @@ -24092,19 +24101,19 @@ x_1 = lean_mk_string_from_bytes("namedArgument", 13); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__116() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__117() { _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_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__2; x_3 = l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__3; -x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__115; +x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__116; 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___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__117() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__118() { _start: { lean_object* x_1; @@ -24112,26 +24121,26 @@ x_1 = lean_mk_string_from_bytes("canonical", 9); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__118() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__119() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__117; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__118; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__119() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__120() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__117; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__118; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__120() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__121() { _start: { lean_object* x_1; lean_object* x_2; @@ -24140,7 +24149,7 @@ x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__121() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__122() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -24150,7 +24159,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__122() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__123() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -24162,19 +24171,19 @@ lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__123() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__124() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__122; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__123; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__124() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__125() { _start: { lean_object* x_1; @@ -24182,26 +24191,26 @@ x_1 = lean_mk_string_from_bytes("r", 1); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__125() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__126() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__124; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__125; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__126() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__127() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__124; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__125; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__127() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__128() { _start: { lean_object* x_1; @@ -24209,82 +24218,82 @@ x_1 = lean_mk_string_from_bytes("mkNullNode", 10); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__128() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__129() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__127; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__128; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__129() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__130() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__127; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__128; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__130() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__131() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__1; -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__127; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__128; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__131() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__132() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__130; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__131; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__132() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__133() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__130; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___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_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__133() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__134() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__132; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__133; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__134() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__135() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__131; -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__133; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__132; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__134; 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___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__135() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__136() { _start: { lean_object* x_1; @@ -24292,17 +24301,17 @@ x_1 = lean_mk_string_from_bytes("term#[_,]", 9); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__136() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__137() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__135; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__136; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__137() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__138() { _start: { lean_object* x_1; @@ -24310,19 +24319,19 @@ x_1 = lean_mk_string_from_bytes("doReturn", 8); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__138() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__139() { _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_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__2; x_3 = l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__3; -x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__137; +x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__138; 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___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__139() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__140() { _start: { lean_object* x_1; @@ -24330,7 +24339,7 @@ x_1 = lean_mk_string_from_bytes("return", 6); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__140() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__141() { _start: { lean_object* x_1; @@ -24338,7 +24347,7 @@ x_1 = lean_mk_string_from_bytes("Termination", 11); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__141() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__142() { _start: { lean_object* x_1; @@ -24346,14 +24355,14 @@ x_1 = lean_mk_string_from_bytes("suffix", 6); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__142() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__143() { _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_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__2; -x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__140; -x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__141; +x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__141; +x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__142; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } @@ -24425,17 +24434,17 @@ x_34 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tac lean_inc_n(x_17, 5); lean_inc(x_12); x_35 = l_Lean_Syntax_node6(x_12, x_34, x_17, x_33, x_17, x_17, x_17, x_17); -x_36 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__15; +x_36 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__17; lean_inc(x_12); x_37 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_37, 0, x_12); lean_ctor_set(x_37, 1, x_36); -x_38 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__19; +x_38 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__20; lean_inc(x_13); lean_inc(x_14); x_39 = l_Lean_addMacroScope(x_14, x_38, x_13); x_40 = lean_box(0); -x_41 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__18; +x_41 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__19; lean_inc(x_12); x_42 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_42, 0, x_12); @@ -24451,24 +24460,24 @@ lean_inc(x_12); x_46 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_46, 0, x_12); lean_ctor_set(x_46, 1, x_45); -x_47 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__26; +x_47 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__27; lean_inc(x_13); lean_inc(x_14); x_48 = l_Lean_addMacroScope(x_14, x_47, x_13); -x_49 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__25; -x_50 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__31; +x_49 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__26; +x_50 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__32; lean_inc(x_12); x_51 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_51, 0, x_12); 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_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__23; +x_52 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__24; lean_inc(x_12); x_53 = l_Lean_Syntax_node2(x_12, x_52, x_46, x_51); lean_inc(x_12); x_54 = l_Lean_Syntax_node1(x_12, x_15, x_53); -x_55 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__21; +x_55 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__22; lean_inc(x_17); lean_inc(x_12); x_56 = l_Lean_Syntax_node2(x_12, x_55, x_17, x_54); @@ -24477,16 +24486,16 @@ lean_inc(x_12); x_58 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_58, 0, x_12); lean_ctor_set(x_58, 1, x_57); -x_59 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__34; +x_59 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__35; lean_inc(x_12); x_60 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_60, 0, x_12); lean_ctor_set(x_60, 1, x_59); -x_61 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__40; +x_61 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__41; lean_inc(x_13); lean_inc(x_14); x_62 = l_Lean_addMacroScope(x_14, x_61, x_13); -x_63 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__39; +x_63 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__40; lean_inc(x_12); x_64 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_64, 0, x_12); @@ -24496,38 +24505,38 @@ lean_ctor_set(x_64, 3, x_40); lean_inc(x_64); lean_inc(x_12); x_65 = l_Lean_Syntax_node1(x_12, x_15, x_64); -x_66 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__41; +x_66 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__42; lean_inc(x_12); x_67 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_67, 0, x_12); lean_ctor_set(x_67, 1, x_66); -x_68 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__42; +x_68 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__43; lean_inc(x_12); x_69 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_69, 0, x_12); lean_ctor_set(x_69, 1, x_68); -x_70 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__50; +x_70 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__51; lean_inc(x_12); x_71 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_71, 0, x_12); lean_ctor_set(x_71, 1, x_70); -x_72 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__55; +x_72 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__56; lean_inc(x_13); lean_inc(x_14); x_73 = l_Lean_addMacroScope(x_14, x_72, x_13); -x_74 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__54; +x_74 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__55; lean_inc(x_12); x_75 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_75, 0, x_12); lean_ctor_set(x_75, 1, x_74); lean_ctor_set(x_75, 2, x_73); lean_ctor_set(x_75, 3, x_40); -x_76 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__56; +x_76 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__57; lean_inc(x_12); x_77 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_77, 0, x_12); lean_ctor_set(x_77, 1, x_76); -x_78 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__59; +x_78 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__60; lean_inc(x_12); x_79 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_79, 0, x_12); @@ -24537,7 +24546,7 @@ lean_inc(x_12); x_81 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_81, 0, x_12); lean_ctor_set(x_81, 1, x_80); -x_82 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__64; +x_82 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__65; lean_inc(x_12); x_83 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_83, 0, x_12); @@ -24545,14 +24554,14 @@ lean_ctor_set(x_83, 1, x_82); x_84 = l_Lean_Syntax_mkNumLit___closed__2; lean_inc(x_12); x_85 = l_Lean_Syntax_node1(x_12, x_84, x_83); -x_86 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__63; +x_86 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__64; lean_inc(x_30); lean_inc(x_85); lean_inc(x_81); lean_inc(x_64); lean_inc(x_12); x_87 = l_Lean_Syntax_node4(x_12, x_86, x_64, x_81, x_85, x_30); -x_88 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__65; +x_88 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__66; lean_inc(x_12); x_89 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_89, 0, x_12); @@ -24564,33 +24573,33 @@ lean_inc(x_90); lean_inc(x_81); lean_inc(x_12); x_91 = l_Lean_Syntax_node4(x_12, x_86, x_87, x_81, x_90, x_30); -x_92 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__61; +x_92 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__62; lean_inc(x_17); lean_inc(x_12); x_93 = l_Lean_Syntax_node2(x_12, x_92, x_17, x_91); lean_inc(x_12); x_94 = l_Lean_Syntax_node1(x_12, x_15, x_93); -x_95 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__66; +x_95 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__67; lean_inc(x_12); x_96 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_96, 0, x_12); lean_ctor_set(x_96, 1, x_95); -x_97 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__71; +x_97 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__72; lean_inc(x_12); x_98 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_98, 0, x_12); lean_ctor_set(x_98, 1, x_97); -x_99 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__74; +x_99 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__75; lean_inc(x_12); x_100 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_100, 0, x_12); lean_ctor_set(x_100, 1, x_99); -x_101 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__76; +x_101 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__77; lean_inc(x_13); lean_inc(x_14); x_102 = l_Lean_addMacroScope(x_14, x_101, x_13); -x_103 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__75; -x_104 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__80; +x_103 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__76; +x_104 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__81; lean_inc(x_12); x_105 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_105, 0, x_12); @@ -24607,12 +24616,12 @@ lean_inc(x_12); x_109 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_109, 0, x_12); lean_ctor_set(x_109, 1, x_108); -x_110 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__84; +x_110 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__85; lean_inc(x_12); x_111 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_111, 0, x_12); lean_ctor_set(x_111, 1, x_110); -x_112 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__83; +x_112 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__84; lean_inc(x_75); lean_inc_n(x_17, 2); lean_inc(x_12); @@ -24630,7 +24639,7 @@ lean_inc(x_109); lean_inc(x_107); lean_inc(x_12); x_117 = l_Lean_Syntax_node5(x_12, x_116, x_107, x_109, x_58, x_113, x_115); -x_118 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__73; +x_118 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__74; lean_inc(x_115); lean_inc(x_98); lean_inc(x_105); @@ -24659,19 +24668,19 @@ lean_inc(x_105); lean_inc(x_100); lean_inc(x_12); x_126 = l_Lean_Syntax_node5(x_12, x_118, x_100, x_105, x_98, x_125, x_115); -x_127 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__86; +x_127 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__87; lean_inc(x_12); x_128 = l_Lean_Syntax_node1(x_12, x_127, x_126); -x_129 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__47; +x_129 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__48; lean_inc(x_17); lean_inc(x_12); x_130 = l_Lean_Syntax_node2(x_12, x_129, x_128, x_17); lean_inc(x_12); x_131 = l_Lean_Syntax_node1(x_12, x_15, x_130); -x_132 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__45; +x_132 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__46; lean_inc(x_12); x_133 = l_Lean_Syntax_node1(x_12, x_132, x_131); -x_134 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__70; +x_134 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__71; lean_inc(x_67); lean_inc(x_98); lean_inc(x_12); @@ -24698,7 +24707,7 @@ lean_inc(x_12); x_145 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_145, 0, x_12); lean_ctor_set(x_145, 1, x_144); -x_146 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__90; +x_146 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__91; lean_inc(x_145); lean_inc(x_143); lean_inc(x_12); @@ -24711,7 +24720,7 @@ x_150 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Ta lean_inc_n(x_17, 3); lean_inc(x_12); x_151 = l_Lean_Syntax_node6(x_12, x_150, x_143, x_17, x_17, x_149, x_17, x_145); -x_152 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__88; +x_152 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__89; lean_inc(x_12); x_153 = l_Lean_Syntax_node2(x_12, x_152, x_147, x_151); lean_inc(x_12); @@ -24741,19 +24750,19 @@ lean_inc(x_12); x_162 = l_Lean_Syntax_node4(x_12, x_134, x_98, x_141, x_67, x_161); lean_inc(x_12); x_163 = l_Lean_Syntax_node2(x_12, x_15, x_135, x_162); -x_164 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__68; +x_164 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__69; lean_inc(x_12); x_165 = l_Lean_Syntax_node1(x_12, x_164, x_163); -x_166 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__58; +x_166 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__59; lean_inc_n(x_17, 2); lean_inc(x_12); x_167 = l_Lean_Syntax_node6(x_12, x_166, x_79, x_17, x_17, x_94, x_96, x_165); -x_168 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__52; +x_168 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__53; lean_inc(x_17); lean_inc(x_75); lean_inc(x_12); x_169 = l_Lean_Syntax_node4(x_12, x_168, x_75, x_17, x_77, x_167); -x_170 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__49; +x_170 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__50; lean_inc(x_17); lean_inc(x_71); lean_inc(x_12); @@ -24761,11 +24770,11 @@ x_171 = l_Lean_Syntax_node3(x_12, x_170, x_71, x_17, x_169); lean_inc(x_17); lean_inc(x_12); x_172 = l_Lean_Syntax_node2(x_12, x_129, x_171, x_17); -x_173 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__100; +x_173 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__101; lean_inc(x_13); lean_inc(x_14); x_174 = l_Lean_addMacroScope(x_14, x_173, x_13); -x_175 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__98; +x_175 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__99; lean_inc(x_12); x_176 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_176, 0, x_12); @@ -24776,16 +24785,16 @@ lean_inc(x_12); x_177 = l_Lean_Syntax_node1(x_12, x_15, x_7); lean_inc(x_12); x_178 = l_Lean_Syntax_node2(x_12, x_123, x_176, x_177); -x_179 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__96; +x_179 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__97; lean_inc(x_58); lean_inc_n(x_17, 2); lean_inc(x_64); lean_inc(x_12); x_180 = l_Lean_Syntax_node5(x_12, x_179, x_64, x_17, x_17, x_58, x_178); -x_181 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__94; +x_181 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__95; lean_inc(x_12); x_182 = l_Lean_Syntax_node1(x_12, x_181, x_180); -x_183 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__92; +x_183 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__93; lean_inc(x_17); lean_inc(x_71); lean_inc(x_12); @@ -24793,23 +24802,23 @@ x_184 = l_Lean_Syntax_node3(x_12, x_183, x_71, x_17, x_182); lean_inc(x_17); lean_inc(x_12); x_185 = l_Lean_Syntax_node2(x_12, x_129, x_184, x_17); -x_186 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__104; +x_186 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__105; lean_inc(x_13); lean_inc(x_14); x_187 = l_Lean_addMacroScope(x_14, x_186, x_13); -x_188 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__102; +x_188 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__103; lean_inc(x_12); x_189 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_189, 0, x_12); lean_ctor_set(x_189, 1, x_188); lean_ctor_set(x_189, 2, x_187); lean_ctor_set(x_189, 3, x_40); -x_190 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__109; +x_190 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__110; lean_inc(x_13); lean_inc(x_14); x_191 = l_Lean_addMacroScope(x_14, x_190, x_13); -x_192 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__108; -x_193 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__114; +x_192 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__109; +x_193 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__115; lean_inc(x_12); x_194 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_194, 0, x_12); @@ -24821,30 +24830,30 @@ lean_inc(x_90); lean_inc(x_64); lean_inc(x_12); x_195 = l_Lean_Syntax_node4(x_12, x_86, x_64, x_81, x_90, x_30); -x_196 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__119; +x_196 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__120; lean_inc(x_13); lean_inc(x_14); x_197 = l_Lean_addMacroScope(x_14, x_196, x_13); -x_198 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__118; +x_198 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__119; lean_inc(x_12); x_199 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_199, 0, x_12); lean_ctor_set(x_199, 1, x_198); lean_ctor_set(x_199, 2, x_197); lean_ctor_set(x_199, 3, x_40); -x_200 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__121; +x_200 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__122; lean_inc(x_13); lean_inc(x_14); x_201 = l_Lean_addMacroScope(x_14, x_200, x_13); -x_202 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__120; -x_203 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__123; +x_202 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__121; +x_203 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__124; lean_inc(x_12); x_204 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_204, 0, x_12); lean_ctor_set(x_204, 1, x_202); lean_ctor_set(x_204, 2, x_201); lean_ctor_set(x_204, 3, x_203); -x_205 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__116; +x_205 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__117; lean_inc(x_115); lean_inc(x_58); lean_inc(x_107); @@ -24854,7 +24863,7 @@ lean_inc(x_12); x_207 = l_Lean_Syntax_node3(x_12, x_15, x_195, x_8, x_206); lean_inc(x_12); x_208 = l_Lean_Syntax_node2(x_12, x_123, x_194, x_207); -x_209 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__106; +x_209 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__107; lean_inc(x_115); lean_inc(x_107); lean_inc(x_12); @@ -24877,21 +24886,21 @@ x_215 = l_Lean_Syntax_node3(x_12, x_183, x_71, x_17, x_214); lean_inc(x_17); lean_inc(x_12); x_216 = l_Lean_Syntax_node2(x_12, x_129, x_215, x_17); -x_217 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__126; +x_217 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__127; lean_inc(x_13); lean_inc(x_14); x_218 = l_Lean_addMacroScope(x_14, x_217, x_13); -x_219 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__125; +x_219 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__126; lean_inc(x_12); x_220 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_220, 0, x_12); lean_ctor_set(x_220, 1, x_219); lean_ctor_set(x_220, 2, x_218); lean_ctor_set(x_220, 3, x_40); -x_221 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__129; +x_221 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__130; x_222 = l_Lean_addMacroScope(x_14, x_221, x_13); -x_223 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__128; -x_224 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__134; +x_223 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__129; +x_224 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__135; lean_inc(x_12); x_225 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_225, 0, x_12); @@ -24905,7 +24914,7 @@ lean_ctor_set(x_227, 0, x_12); lean_ctor_set(x_227, 1, x_226); lean_inc(x_12); x_228 = l_Lean_Syntax_node1(x_12, x_15, x_75); -x_229 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__136; +x_229 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__137; lean_inc(x_12); x_230 = l_Lean_Syntax_node3(x_12, x_229, x_227, x_228, x_30); lean_inc(x_12); @@ -24931,14 +24940,14 @@ x_238 = l_Lean_Syntax_node3(x_12, x_183, x_71, x_17, x_237); lean_inc(x_17); lean_inc(x_12); x_239 = l_Lean_Syntax_node2(x_12, x_129, x_238, x_17); -x_240 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__139; +x_240 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__140; lean_inc(x_12); x_241 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_241, 0, x_12); lean_ctor_set(x_241, 1, x_240); lean_inc(x_12); x_242 = l_Lean_Syntax_node1(x_12, x_15, x_220); -x_243 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__138; +x_243 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__139; lean_inc(x_12); x_244 = l_Lean_Syntax_node2(x_12, x_243, x_241, x_242); lean_inc(x_17); @@ -24948,21 +24957,21 @@ lean_inc(x_12); x_246 = l_Lean_Syntax_node5(x_12, x_15, x_172, x_185, x_216, x_239, x_245); lean_inc(x_12); x_247 = l_Lean_Syntax_node1(x_12, x_132, x_246); -x_248 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__43; +x_248 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__44; lean_inc(x_12); x_249 = l_Lean_Syntax_node2(x_12, x_248, x_69, x_247); -x_250 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__37; +x_250 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__38; lean_inc(x_17); lean_inc(x_12); x_251 = l_Lean_Syntax_node4(x_12, x_250, x_65, x_17, x_67, x_249); -x_252 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__35; +x_252 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__36; lean_inc(x_12); x_253 = l_Lean_Syntax_node2(x_12, x_252, x_60, x_251); -x_254 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__142; +x_254 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__143; lean_inc_n(x_17, 2); lean_inc(x_12); x_255 = l_Lean_Syntax_node2(x_12, x_254, x_17, x_17); -x_256 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__33; +x_256 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__34; lean_inc(x_17); lean_inc(x_12); x_257 = l_Lean_Syntax_node4(x_12, x_256, x_58, x_253, x_255, x_17); @@ -25169,7 +25178,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___aux__Init__Meta______macroRules__Lean__Parser__Syntax__addPrec__1___closed__1; -x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__105; +x_4 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__106; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } @@ -26124,12 +26133,12 @@ x_62 = l_Lean_Syntax_node1(x_28, x_45, x_61); x_63 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__16; lean_inc(x_28); x_64 = l_Lean_Syntax_node1(x_28, x_63, x_15); -x_65 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__76; +x_65 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__77; lean_inc(x_29); lean_inc(x_30); x_66 = l_Lean_addMacroScope(x_30, x_65, x_29); -x_67 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__75; -x_68 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__80; +x_67 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__76; +x_68 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__81; lean_inc(x_28); x_69 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_69, 0, x_28); @@ -26519,12 +26528,12 @@ x_229 = l_Lean_Syntax_node1(x_195, x_212, x_228); x_230 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__16; lean_inc(x_195); x_231 = l_Lean_Syntax_node1(x_195, x_230, x_15); -x_232 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__76; +x_232 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__77; lean_inc(x_196); lean_inc(x_197); x_233 = l_Lean_addMacroScope(x_197, x_232, x_196); -x_234 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__75; -x_235 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__80; +x_234 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__76; +x_235 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__81; lean_inc(x_195); x_236 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_236, 0, x_195); @@ -26892,12 +26901,12 @@ x_387 = l_Lean_Syntax_node1(x_353, x_370, x_386); x_388 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__16; lean_inc(x_353); x_389 = l_Lean_Syntax_node1(x_353, x_388, x_15); -x_390 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__76; +x_390 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__77; lean_inc(x_354); lean_inc(x_355); x_391 = l_Lean_addMacroScope(x_355, x_390, x_354); -x_392 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__75; -x_393 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__80; +x_392 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__76; +x_393 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__81; lean_inc(x_353); x_394 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_394, 0, x_353); @@ -27725,17 +27734,17 @@ lean_inc(x_12); x_20 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_20, 0, x_12); lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__34; +x_21 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__35; lean_inc(x_12); x_22 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_22, 0, x_12); lean_ctor_set(x_22, 1, x_21); -x_23 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__55; +x_23 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__56; lean_inc(x_13); lean_inc(x_14); x_24 = l_Lean_addMacroScope(x_14, x_23, x_13); x_25 = lean_box(0); -x_26 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__54; +x_26 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__55; lean_inc(x_12); x_27 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_27, 0, x_12); @@ -27781,7 +27790,7 @@ x_43 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_43, 0, x_12); lean_ctor_set(x_43, 1, x_35); lean_ctor_set(x_43, 2, x_42); -x_44 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__41; +x_44 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__42; lean_inc(x_12); x_45 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_45, 0, x_12); @@ -27793,7 +27802,7 @@ lean_ctor_set(x_47, 0, x_12); lean_ctor_set(x_47, 1, x_46); lean_inc(x_12); x_48 = l_Lean_Syntax_node1(x_12, x_35, x_27); -x_49 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__66; +x_49 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__67; lean_inc(x_12); x_50 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_50, 0, x_12); @@ -27815,10 +27824,10 @@ x_56 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tac lean_inc(x_43); lean_inc(x_12); x_57 = l_Lean_Syntax_node2(x_12, x_56, x_55, x_43); -x_58 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__121; +x_58 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__122; x_59 = l_Lean_addMacroScope(x_14, x_58, x_13); -x_60 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__120; -x_61 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__123; +x_60 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__121; +x_61 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__124; lean_inc(x_12); x_62 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_62, 0, x_12); @@ -27847,14 +27856,14 @@ lean_inc(x_67); lean_inc(x_47); lean_inc(x_12); x_71 = l_Lean_Syntax_node6(x_12, x_70, x_47, x_51, x_65, x_67, x_43, x_69); -x_72 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__37; +x_72 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__38; lean_inc(x_43); lean_inc(x_12); x_73 = l_Lean_Syntax_node4(x_12, x_72, x_41, x_43, x_45, x_71); -x_74 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__35; +x_74 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__36; lean_inc(x_12); x_75 = l_Lean_Syntax_node2(x_12, x_74, x_22, x_73); -x_76 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__90; +x_76 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__91; lean_inc(x_69); lean_inc(x_47); lean_inc(x_12); @@ -27862,7 +27871,7 @@ x_77 = l_Lean_Syntax_node2(x_12, x_76, x_47, x_69); lean_inc_n(x_43, 2); lean_inc(x_12); x_78 = l_Lean_Syntax_node6(x_12, x_70, x_47, x_43, x_43, x_67, x_43, x_69); -x_79 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__88; +x_79 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__89; lean_inc(x_12); x_80 = l_Lean_Syntax_node2(x_12, x_79, x_77, x_78); lean_inc(x_12); @@ -27904,17 +27913,17 @@ lean_inc(x_90); x_98 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_98, 0, x_90); lean_ctor_set(x_98, 1, x_97); -x_99 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__34; +x_99 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__35; lean_inc(x_90); x_100 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_100, 0, x_90); lean_ctor_set(x_100, 1, x_99); -x_101 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__55; +x_101 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__56; lean_inc(x_91); lean_inc(x_92); x_102 = l_Lean_addMacroScope(x_92, x_101, x_91); x_103 = lean_box(0); -x_104 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__54; +x_104 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__55; lean_inc(x_90); x_105 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_105, 0, x_90); @@ -27960,7 +27969,7 @@ x_121 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_121, 0, x_90); lean_ctor_set(x_121, 1, x_113); lean_ctor_set(x_121, 2, x_120); -x_122 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__41; +x_122 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__42; lean_inc(x_90); x_123 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_123, 0, x_90); @@ -27972,7 +27981,7 @@ lean_ctor_set(x_125, 0, x_90); lean_ctor_set(x_125, 1, x_124); lean_inc(x_90); x_126 = l_Lean_Syntax_node1(x_90, x_113, x_105); -x_127 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__66; +x_127 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__67; lean_inc(x_90); x_128 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_128, 0, x_90); @@ -27994,10 +28003,10 @@ x_134 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Ta lean_inc(x_121); lean_inc(x_90); x_135 = l_Lean_Syntax_node2(x_90, x_134, x_133, x_121); -x_136 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__121; +x_136 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__122; x_137 = l_Lean_addMacroScope(x_92, x_136, x_91); -x_138 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__120; -x_139 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__123; +x_138 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__121; +x_139 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__124; lean_inc(x_90); x_140 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_140, 0, x_90); @@ -28023,10 +28032,10 @@ x_148 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Ta lean_inc(x_121); lean_inc(x_90); x_149 = l_Lean_Syntax_node6(x_90, x_148, x_125, x_129, x_143, x_145, x_121, x_147); -x_150 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__37; +x_150 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__38; lean_inc(x_90); x_151 = l_Lean_Syntax_node4(x_90, x_150, x_119, x_121, x_123, x_149); -x_152 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__35; +x_152 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__36; lean_inc(x_90); x_153 = l_Lean_Syntax_node2(x_90, x_152, x_100, x_151); lean_inc(x_90); @@ -28301,17 +28310,17 @@ lean_inc(x_12); x_20 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_20, 0, x_12); lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__34; +x_21 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__35; lean_inc(x_12); x_22 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_22, 0, x_12); lean_ctor_set(x_22, 1, x_21); -x_23 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__55; +x_23 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__56; lean_inc(x_13); lean_inc(x_14); x_24 = l_Lean_addMacroScope(x_14, x_23, x_13); x_25 = lean_box(0); -x_26 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__54; +x_26 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__55; lean_inc(x_12); x_27 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_27, 0, x_12); @@ -28357,7 +28366,7 @@ x_43 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_43, 0, x_12); lean_ctor_set(x_43, 1, x_35); lean_ctor_set(x_43, 2, x_42); -x_44 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__41; +x_44 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__42; lean_inc(x_12); x_45 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_45, 0, x_12); @@ -28369,7 +28378,7 @@ lean_ctor_set(x_47, 0, x_12); lean_ctor_set(x_47, 1, x_46); lean_inc(x_12); x_48 = l_Lean_Syntax_node1(x_12, x_35, x_27); -x_49 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__66; +x_49 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__67; lean_inc(x_12); x_50 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_50, 0, x_12); @@ -28391,12 +28400,12 @@ x_56 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tac lean_inc(x_43); lean_inc(x_12); x_57 = l_Lean_Syntax_node2(x_12, x_56, x_55, x_43); -x_58 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__121; +x_58 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__122; lean_inc(x_13); lean_inc(x_14); x_59 = l_Lean_addMacroScope(x_14, x_58, x_13); -x_60 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__120; -x_61 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__123; +x_60 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__121; +x_61 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__124; lean_inc(x_12); x_62 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_62, 0, x_12); @@ -28447,14 +28456,14 @@ lean_inc(x_76); lean_inc(x_47); lean_inc(x_12); x_80 = l_Lean_Syntax_node6(x_12, x_79, x_47, x_51, x_74, x_76, x_43, x_78); -x_81 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__37; +x_81 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__38; lean_inc(x_43); lean_inc(x_12); x_82 = l_Lean_Syntax_node4(x_12, x_81, x_41, x_43, x_45, x_80); -x_83 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__35; +x_83 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__36; lean_inc(x_12); x_84 = l_Lean_Syntax_node2(x_12, x_83, x_22, x_82); -x_85 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__90; +x_85 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__91; lean_inc(x_78); lean_inc(x_47); lean_inc(x_12); @@ -28462,7 +28471,7 @@ x_86 = l_Lean_Syntax_node2(x_12, x_85, x_47, x_78); lean_inc_n(x_43, 2); lean_inc(x_12); x_87 = l_Lean_Syntax_node6(x_12, x_79, x_47, x_43, x_43, x_76, x_43, x_78); -x_88 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__88; +x_88 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__89; lean_inc(x_12); x_89 = l_Lean_Syntax_node2(x_12, x_88, x_86, x_87); lean_inc(x_12); @@ -28504,17 +28513,17 @@ lean_inc(x_99); x_107 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_107, 0, x_99); lean_ctor_set(x_107, 1, x_106); -x_108 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__34; +x_108 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__35; lean_inc(x_99); x_109 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_109, 0, x_99); lean_ctor_set(x_109, 1, x_108); -x_110 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__55; +x_110 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__56; lean_inc(x_100); lean_inc(x_101); x_111 = l_Lean_addMacroScope(x_101, x_110, x_100); x_112 = lean_box(0); -x_113 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__54; +x_113 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__55; lean_inc(x_99); x_114 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_114, 0, x_99); @@ -28560,7 +28569,7 @@ x_130 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_130, 0, x_99); lean_ctor_set(x_130, 1, x_122); lean_ctor_set(x_130, 2, x_129); -x_131 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__41; +x_131 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__42; lean_inc(x_99); x_132 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_132, 0, x_99); @@ -28572,7 +28581,7 @@ lean_ctor_set(x_134, 0, x_99); lean_ctor_set(x_134, 1, x_133); lean_inc(x_99); x_135 = l_Lean_Syntax_node1(x_99, x_122, x_114); -x_136 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__66; +x_136 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__67; lean_inc(x_99); x_137 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_137, 0, x_99); @@ -28594,12 +28603,12 @@ x_143 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Ta lean_inc(x_130); lean_inc(x_99); x_144 = l_Lean_Syntax_node2(x_99, x_143, x_142, x_130); -x_145 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__121; +x_145 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__122; lean_inc(x_100); lean_inc(x_101); x_146 = l_Lean_addMacroScope(x_101, x_145, x_100); -x_147 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__120; -x_148 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__123; +x_147 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__121; +x_148 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__124; lean_inc(x_99); x_149 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_149, 0, x_99); @@ -28647,10 +28656,10 @@ x_166 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Ta lean_inc(x_130); lean_inc(x_99); x_167 = l_Lean_Syntax_node6(x_99, x_166, x_134, x_138, x_161, x_163, x_130, x_165); -x_168 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__37; +x_168 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__38; lean_inc(x_99); x_169 = l_Lean_Syntax_node4(x_99, x_168, x_128, x_130, x_132, x_167); -x_170 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__35; +x_170 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__36; lean_inc(x_99); x_171 = l_Lean_Syntax_node2(x_99, x_170, x_109, x_169); lean_inc(x_99); @@ -28836,17 +28845,17 @@ lean_inc(x_12); x_20 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_20, 0, x_12); lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__34; +x_21 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__35; lean_inc(x_12); x_22 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_22, 0, x_12); lean_ctor_set(x_22, 1, x_21); -x_23 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__55; +x_23 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__56; lean_inc(x_13); lean_inc(x_14); x_24 = l_Lean_addMacroScope(x_14, x_23, x_13); x_25 = lean_box(0); -x_26 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__54; +x_26 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__55; lean_inc(x_12); x_27 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_27, 0, x_12); @@ -28892,7 +28901,7 @@ x_43 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_43, 0, x_12); lean_ctor_set(x_43, 1, x_35); lean_ctor_set(x_43, 2, x_42); -x_44 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__41; +x_44 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__42; lean_inc(x_12); x_45 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_45, 0, x_12); @@ -28904,7 +28913,7 @@ lean_ctor_set(x_47, 0, x_12); lean_ctor_set(x_47, 1, x_46); lean_inc(x_12); x_48 = l_Lean_Syntax_node1(x_12, x_35, x_27); -x_49 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__66; +x_49 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__67; lean_inc(x_12); x_50 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_50, 0, x_12); @@ -28926,12 +28935,12 @@ x_56 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tac lean_inc(x_43); lean_inc(x_12); x_57 = l_Lean_Syntax_node2(x_12, x_56, x_55, x_43); -x_58 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__121; +x_58 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__122; lean_inc(x_13); lean_inc(x_14); x_59 = l_Lean_addMacroScope(x_14, x_58, x_13); -x_60 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__120; -x_61 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__123; +x_60 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__121; +x_61 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__124; lean_inc(x_12); x_62 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_62, 0, x_12); @@ -29001,14 +29010,14 @@ lean_inc(x_82); lean_inc(x_47); lean_inc(x_12); x_86 = l_Lean_Syntax_node6(x_12, x_85, x_47, x_51, x_80, x_82, x_43, x_84); -x_87 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__37; +x_87 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__38; lean_inc(x_43); lean_inc(x_12); x_88 = l_Lean_Syntax_node4(x_12, x_87, x_41, x_43, x_45, x_86); -x_89 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__35; +x_89 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__36; lean_inc(x_12); x_90 = l_Lean_Syntax_node2(x_12, x_89, x_22, x_88); -x_91 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__90; +x_91 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__91; lean_inc(x_84); lean_inc(x_47); lean_inc(x_12); @@ -29016,7 +29025,7 @@ x_92 = l_Lean_Syntax_node2(x_12, x_91, x_47, x_84); lean_inc_n(x_43, 2); lean_inc(x_12); x_93 = l_Lean_Syntax_node6(x_12, x_85, x_47, x_43, x_43, x_82, x_43, x_84); -x_94 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__88; +x_94 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__89; lean_inc(x_12); x_95 = l_Lean_Syntax_node2(x_12, x_94, x_92, x_93); lean_inc(x_12); @@ -29058,17 +29067,17 @@ lean_inc(x_105); x_113 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_113, 0, x_105); lean_ctor_set(x_113, 1, x_112); -x_114 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__34; +x_114 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__35; lean_inc(x_105); x_115 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_115, 0, x_105); lean_ctor_set(x_115, 1, x_114); -x_116 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__55; +x_116 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__56; lean_inc(x_106); lean_inc(x_107); x_117 = l_Lean_addMacroScope(x_107, x_116, x_106); x_118 = lean_box(0); -x_119 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__54; +x_119 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__55; lean_inc(x_105); x_120 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_120, 0, x_105); @@ -29114,7 +29123,7 @@ x_136 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_136, 0, x_105); lean_ctor_set(x_136, 1, x_128); lean_ctor_set(x_136, 2, x_135); -x_137 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__41; +x_137 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__42; lean_inc(x_105); x_138 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_138, 0, x_105); @@ -29126,7 +29135,7 @@ lean_ctor_set(x_140, 0, x_105); lean_ctor_set(x_140, 1, x_139); lean_inc(x_105); x_141 = l_Lean_Syntax_node1(x_105, x_128, x_120); -x_142 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__66; +x_142 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__67; lean_inc(x_105); x_143 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_143, 0, x_105); @@ -29148,12 +29157,12 @@ x_149 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Ta lean_inc(x_136); lean_inc(x_105); x_150 = l_Lean_Syntax_node2(x_105, x_149, x_148, x_136); -x_151 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__121; +x_151 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__122; lean_inc(x_106); lean_inc(x_107); x_152 = l_Lean_addMacroScope(x_107, x_151, x_106); -x_153 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__120; -x_154 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__123; +x_153 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__121; +x_154 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__124; lean_inc(x_105); x_155 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_155, 0, x_105); @@ -29220,10 +29229,10 @@ x_178 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Ta lean_inc(x_136); lean_inc(x_105); x_179 = l_Lean_Syntax_node6(x_105, x_178, x_140, x_144, x_173, x_175, x_136, x_177); -x_180 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__37; +x_180 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__38; lean_inc(x_105); x_181 = l_Lean_Syntax_node4(x_105, x_180, x_134, x_136, x_138, x_179); -x_182 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__35; +x_182 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__36; lean_inc(x_105); x_183 = l_Lean_Syntax_node2(x_105, x_182, x_115, x_181); lean_inc(x_105); @@ -29570,17 +29579,17 @@ lean_inc(x_12); x_20 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_20, 0, x_12); lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__34; +x_21 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__35; lean_inc(x_12); x_22 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_22, 0, x_12); lean_ctor_set(x_22, 1, x_21); -x_23 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__55; +x_23 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__56; lean_inc(x_13); lean_inc(x_14); x_24 = l_Lean_addMacroScope(x_14, x_23, x_13); x_25 = lean_box(0); -x_26 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__54; +x_26 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__55; lean_inc(x_12); x_27 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_27, 0, x_12); @@ -29626,7 +29635,7 @@ x_43 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_43, 0, x_12); lean_ctor_set(x_43, 1, x_35); lean_ctor_set(x_43, 2, x_42); -x_44 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__41; +x_44 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__42; lean_inc(x_12); x_45 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_45, 0, x_12); @@ -29638,7 +29647,7 @@ lean_ctor_set(x_47, 0, x_12); lean_ctor_set(x_47, 1, x_46); lean_inc(x_12); x_48 = l_Lean_Syntax_node1(x_12, x_35, x_27); -x_49 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__66; +x_49 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__67; lean_inc(x_12); x_50 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_50, 0, x_12); @@ -29660,10 +29669,10 @@ x_56 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tac lean_inc(x_43); lean_inc(x_12); x_57 = l_Lean_Syntax_node2(x_12, x_56, x_55, x_43); -x_58 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__121; +x_58 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__122; x_59 = l_Lean_addMacroScope(x_14, x_58, x_13); -x_60 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__120; -x_61 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__123; +x_60 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__121; +x_61 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__124; lean_inc(x_12); x_62 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_62, 0, x_12); @@ -29692,14 +29701,14 @@ lean_inc(x_67); lean_inc(x_47); lean_inc(x_12); x_71 = l_Lean_Syntax_node6(x_12, x_70, x_47, x_51, x_65, x_67, x_43, x_69); -x_72 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__37; +x_72 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__38; lean_inc(x_43); lean_inc(x_12); x_73 = l_Lean_Syntax_node4(x_12, x_72, x_41, x_43, x_45, x_71); -x_74 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__35; +x_74 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__36; lean_inc(x_12); x_75 = l_Lean_Syntax_node2(x_12, x_74, x_22, x_73); -x_76 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__90; +x_76 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__91; lean_inc(x_69); lean_inc(x_47); lean_inc(x_12); @@ -29707,7 +29716,7 @@ x_77 = l_Lean_Syntax_node2(x_12, x_76, x_47, x_69); lean_inc_n(x_43, 2); lean_inc(x_12); x_78 = l_Lean_Syntax_node6(x_12, x_70, x_47, x_43, x_43, x_67, x_43, x_69); -x_79 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__88; +x_79 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__89; lean_inc(x_12); x_80 = l_Lean_Syntax_node2(x_12, x_79, x_77, x_78); lean_inc(x_12); @@ -29749,17 +29758,17 @@ lean_inc(x_90); x_98 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_98, 0, x_90); lean_ctor_set(x_98, 1, x_97); -x_99 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__34; +x_99 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__35; lean_inc(x_90); x_100 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_100, 0, x_90); lean_ctor_set(x_100, 1, x_99); -x_101 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__55; +x_101 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__56; lean_inc(x_91); lean_inc(x_92); x_102 = l_Lean_addMacroScope(x_92, x_101, x_91); x_103 = lean_box(0); -x_104 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__54; +x_104 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__55; lean_inc(x_90); x_105 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_105, 0, x_90); @@ -29805,7 +29814,7 @@ x_121 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_121, 0, x_90); lean_ctor_set(x_121, 1, x_113); lean_ctor_set(x_121, 2, x_120); -x_122 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__41; +x_122 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__42; lean_inc(x_90); x_123 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_123, 0, x_90); @@ -29817,7 +29826,7 @@ lean_ctor_set(x_125, 0, x_90); lean_ctor_set(x_125, 1, x_124); lean_inc(x_90); x_126 = l_Lean_Syntax_node1(x_90, x_113, x_105); -x_127 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__66; +x_127 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__67; lean_inc(x_90); x_128 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_128, 0, x_90); @@ -29839,10 +29848,10 @@ x_134 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Ta lean_inc(x_121); lean_inc(x_90); x_135 = l_Lean_Syntax_node2(x_90, x_134, x_133, x_121); -x_136 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__121; +x_136 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__122; x_137 = l_Lean_addMacroScope(x_92, x_136, x_91); -x_138 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__120; -x_139 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__123; +x_138 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__121; +x_139 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__124; lean_inc(x_90); x_140 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_140, 0, x_90); @@ -29868,10 +29877,10 @@ x_148 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Ta lean_inc(x_121); lean_inc(x_90); x_149 = l_Lean_Syntax_node6(x_90, x_148, x_125, x_129, x_143, x_145, x_121, x_147); -x_150 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__37; +x_150 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__38; lean_inc(x_90); x_151 = l_Lean_Syntax_node4(x_90, x_150, x_119, x_121, x_123, x_149); -x_152 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__35; +x_152 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__36; lean_inc(x_90); x_153 = l_Lean_Syntax_node2(x_90, x_152, x_100, x_151); lean_inc(x_90); @@ -30052,17 +30061,17 @@ lean_inc(x_12); x_20 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_20, 0, x_12); lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__34; +x_21 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__35; lean_inc(x_12); x_22 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_22, 0, x_12); lean_ctor_set(x_22, 1, x_21); -x_23 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__55; +x_23 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__56; lean_inc(x_13); lean_inc(x_14); x_24 = l_Lean_addMacroScope(x_14, x_23, x_13); x_25 = lean_box(0); -x_26 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__54; +x_26 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__55; lean_inc(x_12); x_27 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_27, 0, x_12); @@ -30108,7 +30117,7 @@ x_43 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_43, 0, x_12); lean_ctor_set(x_43, 1, x_35); lean_ctor_set(x_43, 2, x_42); -x_44 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__41; +x_44 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__42; lean_inc(x_12); x_45 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_45, 0, x_12); @@ -30120,7 +30129,7 @@ lean_ctor_set(x_47, 0, x_12); lean_ctor_set(x_47, 1, x_46); lean_inc(x_12); x_48 = l_Lean_Syntax_node1(x_12, x_35, x_27); -x_49 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__66; +x_49 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__67; lean_inc(x_12); x_50 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_50, 0, x_12); @@ -30142,12 +30151,12 @@ x_56 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tac lean_inc(x_43); lean_inc(x_12); x_57 = l_Lean_Syntax_node2(x_12, x_56, x_55, x_43); -x_58 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__121; +x_58 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__122; lean_inc(x_13); lean_inc(x_14); x_59 = l_Lean_addMacroScope(x_14, x_58, x_13); -x_60 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__120; -x_61 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__123; +x_60 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__121; +x_61 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__124; lean_inc(x_12); x_62 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_62, 0, x_12); @@ -30198,14 +30207,14 @@ lean_inc(x_76); lean_inc(x_47); lean_inc(x_12); x_80 = l_Lean_Syntax_node6(x_12, x_79, x_47, x_51, x_74, x_76, x_43, x_78); -x_81 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__37; +x_81 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__38; lean_inc(x_43); lean_inc(x_12); x_82 = l_Lean_Syntax_node4(x_12, x_81, x_41, x_43, x_45, x_80); -x_83 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__35; +x_83 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__36; lean_inc(x_12); x_84 = l_Lean_Syntax_node2(x_12, x_83, x_22, x_82); -x_85 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__90; +x_85 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__91; lean_inc(x_78); lean_inc(x_47); lean_inc(x_12); @@ -30213,7 +30222,7 @@ x_86 = l_Lean_Syntax_node2(x_12, x_85, x_47, x_78); lean_inc_n(x_43, 2); lean_inc(x_12); x_87 = l_Lean_Syntax_node6(x_12, x_79, x_47, x_43, x_43, x_76, x_43, x_78); -x_88 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__88; +x_88 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__89; lean_inc(x_12); x_89 = l_Lean_Syntax_node2(x_12, x_88, x_86, x_87); lean_inc(x_12); @@ -30255,17 +30264,17 @@ lean_inc(x_99); x_107 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_107, 0, x_99); lean_ctor_set(x_107, 1, x_106); -x_108 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__34; +x_108 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__35; lean_inc(x_99); x_109 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_109, 0, x_99); lean_ctor_set(x_109, 1, x_108); -x_110 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__55; +x_110 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__56; lean_inc(x_100); lean_inc(x_101); x_111 = l_Lean_addMacroScope(x_101, x_110, x_100); x_112 = lean_box(0); -x_113 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__54; +x_113 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__55; lean_inc(x_99); x_114 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_114, 0, x_99); @@ -30311,7 +30320,7 @@ x_130 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_130, 0, x_99); lean_ctor_set(x_130, 1, x_122); lean_ctor_set(x_130, 2, x_129); -x_131 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__41; +x_131 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__42; lean_inc(x_99); x_132 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_132, 0, x_99); @@ -30323,7 +30332,7 @@ lean_ctor_set(x_134, 0, x_99); lean_ctor_set(x_134, 1, x_133); lean_inc(x_99); x_135 = l_Lean_Syntax_node1(x_99, x_122, x_114); -x_136 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__66; +x_136 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__67; lean_inc(x_99); x_137 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_137, 0, x_99); @@ -30345,12 +30354,12 @@ x_143 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Ta lean_inc(x_130); lean_inc(x_99); x_144 = l_Lean_Syntax_node2(x_99, x_143, x_142, x_130); -x_145 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__121; +x_145 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__122; lean_inc(x_100); lean_inc(x_101); x_146 = l_Lean_addMacroScope(x_101, x_145, x_100); -x_147 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__120; -x_148 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__123; +x_147 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__121; +x_148 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__124; lean_inc(x_99); x_149 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_149, 0, x_99); @@ -30398,10 +30407,10 @@ x_166 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Ta lean_inc(x_130); lean_inc(x_99); x_167 = l_Lean_Syntax_node6(x_99, x_166, x_134, x_138, x_161, x_163, x_130, x_165); -x_168 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__37; +x_168 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__38; lean_inc(x_99); x_169 = l_Lean_Syntax_node4(x_99, x_168, x_128, x_130, x_132, x_167); -x_170 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__35; +x_170 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__36; lean_inc(x_99); x_171 = l_Lean_Syntax_node2(x_99, x_170, x_109, x_169); lean_inc(x_99); @@ -30573,17 +30582,17 @@ lean_inc(x_12); x_20 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_20, 0, x_12); lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__34; +x_21 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__35; lean_inc(x_12); x_22 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_22, 0, x_12); lean_ctor_set(x_22, 1, x_21); -x_23 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__55; +x_23 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__56; lean_inc(x_13); lean_inc(x_14); x_24 = l_Lean_addMacroScope(x_14, x_23, x_13); x_25 = lean_box(0); -x_26 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__54; +x_26 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__55; lean_inc(x_12); x_27 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_27, 0, x_12); @@ -30629,7 +30638,7 @@ x_43 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_43, 0, x_12); lean_ctor_set(x_43, 1, x_35); lean_ctor_set(x_43, 2, x_42); -x_44 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__41; +x_44 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__42; lean_inc(x_12); x_45 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_45, 0, x_12); @@ -30641,7 +30650,7 @@ lean_ctor_set(x_47, 0, x_12); lean_ctor_set(x_47, 1, x_46); lean_inc(x_12); x_48 = l_Lean_Syntax_node1(x_12, x_35, x_27); -x_49 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__66; +x_49 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__67; lean_inc(x_12); x_50 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_50, 0, x_12); @@ -30663,12 +30672,12 @@ x_56 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tac lean_inc(x_43); lean_inc(x_12); x_57 = l_Lean_Syntax_node2(x_12, x_56, x_55, x_43); -x_58 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__121; +x_58 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__122; lean_inc(x_13); lean_inc(x_14); x_59 = l_Lean_addMacroScope(x_14, x_58, x_13); -x_60 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__120; -x_61 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__123; +x_60 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__121; +x_61 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__124; lean_inc(x_12); x_62 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_62, 0, x_12); @@ -30738,14 +30747,14 @@ lean_inc(x_82); lean_inc(x_47); lean_inc(x_12); x_86 = l_Lean_Syntax_node6(x_12, x_85, x_47, x_51, x_80, x_82, x_43, x_84); -x_87 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__37; +x_87 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__38; lean_inc(x_43); lean_inc(x_12); x_88 = l_Lean_Syntax_node4(x_12, x_87, x_41, x_43, x_45, x_86); -x_89 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__35; +x_89 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__36; lean_inc(x_12); x_90 = l_Lean_Syntax_node2(x_12, x_89, x_22, x_88); -x_91 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__90; +x_91 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__91; lean_inc(x_84); lean_inc(x_47); lean_inc(x_12); @@ -30753,7 +30762,7 @@ x_92 = l_Lean_Syntax_node2(x_12, x_91, x_47, x_84); lean_inc_n(x_43, 2); lean_inc(x_12); x_93 = l_Lean_Syntax_node6(x_12, x_85, x_47, x_43, x_43, x_82, x_43, x_84); -x_94 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__88; +x_94 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__89; lean_inc(x_12); x_95 = l_Lean_Syntax_node2(x_12, x_94, x_92, x_93); lean_inc(x_12); @@ -30795,17 +30804,17 @@ lean_inc(x_105); x_113 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_113, 0, x_105); lean_ctor_set(x_113, 1, x_112); -x_114 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__34; +x_114 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__35; lean_inc(x_105); x_115 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_115, 0, x_105); lean_ctor_set(x_115, 1, x_114); -x_116 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__55; +x_116 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__56; lean_inc(x_106); lean_inc(x_107); x_117 = l_Lean_addMacroScope(x_107, x_116, x_106); x_118 = lean_box(0); -x_119 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__54; +x_119 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__55; lean_inc(x_105); x_120 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_120, 0, x_105); @@ -30851,7 +30860,7 @@ x_136 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_136, 0, x_105); lean_ctor_set(x_136, 1, x_128); lean_ctor_set(x_136, 2, x_135); -x_137 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__41; +x_137 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__42; lean_inc(x_105); x_138 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_138, 0, x_105); @@ -30863,7 +30872,7 @@ lean_ctor_set(x_140, 0, x_105); lean_ctor_set(x_140, 1, x_139); lean_inc(x_105); x_141 = l_Lean_Syntax_node1(x_105, x_128, x_120); -x_142 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__66; +x_142 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__67; lean_inc(x_105); x_143 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_143, 0, x_105); @@ -30885,12 +30894,12 @@ x_149 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Ta lean_inc(x_136); lean_inc(x_105); x_150 = l_Lean_Syntax_node2(x_105, x_149, x_148, x_136); -x_151 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__121; +x_151 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__122; lean_inc(x_106); lean_inc(x_107); x_152 = l_Lean_addMacroScope(x_107, x_151, x_106); -x_153 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__120; -x_154 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__123; +x_153 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__121; +x_154 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__124; lean_inc(x_105); x_155 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_155, 0, x_105); @@ -30957,10 +30966,10 @@ x_178 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Ta lean_inc(x_136); lean_inc(x_105); x_179 = l_Lean_Syntax_node6(x_105, x_178, x_140, x_144, x_173, x_175, x_136, x_177); -x_180 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__37; +x_180 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__38; lean_inc(x_105); x_181 = l_Lean_Syntax_node4(x_105, x_180, x_134, x_136, x_138, x_179); -x_182 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__35; +x_182 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__36; lean_inc(x_105); x_183 = l_Lean_Syntax_node2(x_105, x_182, x_115, x_181); lean_inc(x_105); @@ -31257,17 +31266,17 @@ lean_inc(x_12); x_20 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_20, 0, x_12); lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__34; +x_21 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__35; lean_inc(x_12); x_22 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_22, 0, x_12); lean_ctor_set(x_22, 1, x_21); -x_23 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__55; +x_23 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__56; lean_inc(x_13); lean_inc(x_14); x_24 = l_Lean_addMacroScope(x_14, x_23, x_13); x_25 = lean_box(0); -x_26 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__54; +x_26 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__55; lean_inc(x_12); x_27 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_27, 0, x_12); @@ -31313,7 +31322,7 @@ x_43 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_43, 0, x_12); lean_ctor_set(x_43, 1, x_35); lean_ctor_set(x_43, 2, x_42); -x_44 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__41; +x_44 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__42; lean_inc(x_12); x_45 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_45, 0, x_12); @@ -31325,7 +31334,7 @@ lean_ctor_set(x_47, 0, x_12); lean_ctor_set(x_47, 1, x_46); lean_inc(x_12); x_48 = l_Lean_Syntax_node1(x_12, x_35, x_27); -x_49 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__66; +x_49 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__67; lean_inc(x_12); x_50 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_50, 0, x_12); @@ -31347,10 +31356,10 @@ x_56 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tac lean_inc(x_43); lean_inc(x_12); x_57 = l_Lean_Syntax_node2(x_12, x_56, x_55, x_43); -x_58 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__121; +x_58 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__122; x_59 = l_Lean_addMacroScope(x_14, x_58, x_13); -x_60 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__120; -x_61 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__123; +x_60 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__121; +x_61 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__124; lean_inc(x_12); x_62 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_62, 0, x_12); @@ -31379,14 +31388,14 @@ lean_inc(x_67); lean_inc(x_47); lean_inc(x_12); x_71 = l_Lean_Syntax_node6(x_12, x_70, x_47, x_51, x_65, x_67, x_43, x_69); -x_72 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__37; +x_72 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__38; lean_inc(x_43); lean_inc(x_12); x_73 = l_Lean_Syntax_node4(x_12, x_72, x_41, x_43, x_45, x_71); -x_74 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__35; +x_74 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__36; lean_inc(x_12); x_75 = l_Lean_Syntax_node2(x_12, x_74, x_22, x_73); -x_76 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__90; +x_76 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__91; lean_inc(x_69); lean_inc(x_47); lean_inc(x_12); @@ -31394,7 +31403,7 @@ x_77 = l_Lean_Syntax_node2(x_12, x_76, x_47, x_69); lean_inc_n(x_43, 2); lean_inc(x_12); x_78 = l_Lean_Syntax_node6(x_12, x_70, x_47, x_43, x_43, x_67, x_43, x_69); -x_79 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__88; +x_79 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__89; lean_inc(x_12); x_80 = l_Lean_Syntax_node2(x_12, x_79, x_77, x_78); lean_inc(x_12); @@ -31436,17 +31445,17 @@ lean_inc(x_90); x_98 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_98, 0, x_90); lean_ctor_set(x_98, 1, x_97); -x_99 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__34; +x_99 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__35; lean_inc(x_90); x_100 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_100, 0, x_90); lean_ctor_set(x_100, 1, x_99); -x_101 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__55; +x_101 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__56; lean_inc(x_91); lean_inc(x_92); x_102 = l_Lean_addMacroScope(x_92, x_101, x_91); x_103 = lean_box(0); -x_104 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__54; +x_104 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__55; lean_inc(x_90); x_105 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_105, 0, x_90); @@ -31492,7 +31501,7 @@ x_121 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_121, 0, x_90); lean_ctor_set(x_121, 1, x_113); lean_ctor_set(x_121, 2, x_120); -x_122 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__41; +x_122 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__42; lean_inc(x_90); x_123 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_123, 0, x_90); @@ -31504,7 +31513,7 @@ lean_ctor_set(x_125, 0, x_90); lean_ctor_set(x_125, 1, x_124); lean_inc(x_90); x_126 = l_Lean_Syntax_node1(x_90, x_113, x_105); -x_127 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__66; +x_127 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__67; lean_inc(x_90); x_128 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_128, 0, x_90); @@ -31526,10 +31535,10 @@ x_134 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Ta lean_inc(x_121); lean_inc(x_90); x_135 = l_Lean_Syntax_node2(x_90, x_134, x_133, x_121); -x_136 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__121; +x_136 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__122; x_137 = l_Lean_addMacroScope(x_92, x_136, x_91); -x_138 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__120; -x_139 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__123; +x_138 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__121; +x_139 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__124; lean_inc(x_90); x_140 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_140, 0, x_90); @@ -31555,10 +31564,10 @@ x_148 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Ta lean_inc(x_121); lean_inc(x_90); x_149 = l_Lean_Syntax_node6(x_90, x_148, x_125, x_129, x_143, x_145, x_121, x_147); -x_150 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__37; +x_150 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__38; lean_inc(x_90); x_151 = l_Lean_Syntax_node4(x_90, x_150, x_119, x_121, x_123, x_149); -x_152 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__35; +x_152 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__36; lean_inc(x_90); x_153 = l_Lean_Syntax_node2(x_90, x_152, x_100, x_151); lean_inc(x_90); @@ -32867,6 +32876,8 @@ l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__de lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__141); l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__142 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__142(); lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__142); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__143 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__143(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__143); l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__1 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__1(); lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__1); l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__2 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__2(); diff --git a/stage0/stdlib/Init/NotationExtra.c b/stage0/stdlib/Init/NotationExtra.c index 5e28ceaab107..1ff881d4b4a3 100644 --- a/stage0/stdlib/Init/NotationExtra.c +++ b/stage0/stdlib/Init/NotationExtra.c @@ -769,6 +769,7 @@ static lean_object* l_Lean_doElemWhile___x3a__Do_____closed__12; static lean_object* l_Lean_command____Unif__hint________Where___x7c___x2d_u22a2_____closed__22; LEAN_EXPORT lean_object* l_Lean_unbracketedExplicitBinders; static lean_object* l_Lean_doElemWhile__Do_____closed__5; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__49; 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_Lean_doElemRepeat_____closed__7; LEAN_EXPORT lean_object* l_Lean_Loop_toCtorIdx(lean_object*); @@ -4251,7 +4252,7 @@ static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lea _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("def", 3); +x_1 = lean_mk_string_from_bytes("definition", 10); return x_1; } } @@ -4271,23 +4272,31 @@ static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lea _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("declId", 6); +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__30() { _start: { +lean_object* x_1; +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__31() { +_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__11; -x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__29; +x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__30; 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__31() { +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; @@ -4301,7 +4310,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__32() { +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; @@ -4310,7 +4319,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__33() { +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; @@ -4318,19 +4327,19 @@ 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__34() { +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; 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__11; -x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__33; +x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__34; 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__35() { +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; @@ -4342,19 +4351,19 @@ 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__37() { _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__36; 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__38() { _start: { lean_object* x_1; @@ -4362,19 +4371,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__39() { _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__38; 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__40() { _start: { lean_object* x_1; @@ -4382,7 +4391,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__41() { _start: { lean_object* x_1; @@ -4390,19 +4399,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__42() { _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__41; 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__43() { _start: { lean_object* x_1; @@ -4410,19 +4419,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__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_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__11; -x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__42; +x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___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___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__45() { _start: { lean_object* x_1; @@ -4430,7 +4439,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__46() { _start: { lean_object* x_1; @@ -4438,7 +4447,7 @@ x_1 = lean_mk_string_from_bytes("Termination", 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__46() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__47() { _start: { lean_object* x_1; @@ -4446,19 +4455,19 @@ x_1 = lean_mk_string_from_bytes("suffix", 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__47() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__48() { _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__45; -x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__46; +x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__46; +x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__47; 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__48() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__49() { _start: { lean_object* x_1; lean_object* x_2; @@ -4679,7 +4688,7 @@ lean_inc(x_45); x_81 = l_Lean_Syntax_node3(x_45, x_80, x_65, x_77, x_79); lean_inc(x_45); x_82 = l_Lean_Syntax_node1(x_45, x_70, x_81); -x_83 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__27; +x_83 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__29; lean_inc(x_45); x_84 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_84, 0, x_45); @@ -4687,12 +4696,12 @@ lean_ctor_set(x_84, 1, x_83); x_85 = lean_array_get_size(x_40); 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_87 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__36; 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_40); 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_91 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__37; 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_45); @@ -4705,7 +4714,7 @@ lean_inc(x_45); x_96 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_96, 0, x_45); 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__40; lean_inc(x_45); x_98 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_98, 0, x_45); @@ -4715,12 +4724,12 @@ lean_inc(x_45); x_100 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_100, 0, x_45); 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__42; lean_inc(x_45); x_102 = l_Lean_Syntax_node1(x_45, x_101, x_100); lean_inc(x_45); x_103 = l_Lean_Syntax_node1(x_45, 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__39; lean_inc(x_45); x_105 = l_Lean_Syntax_node2(x_45, x_104, x_98, x_103); x_106 = l_Lean_expandExplicitBindersAux_loop___closed__11; @@ -4728,26 +4737,26 @@ lean_inc(x_45); x_107 = l_Lean_Syntax_node2(x_45, x_106, x_96, x_105); lean_inc(x_45); x_108 = l_Lean_Syntax_node1(x_45, x_70, x_107); -x_109 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__34; +x_109 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__35; lean_inc(x_45); x_110 = l_Lean_Syntax_node2(x_45, 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__45; lean_inc(x_45); x_112 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_112, 0, x_45); 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__47; +x_113 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__48; lean_inc_n(x_72, 2); lean_inc(x_45); x_114 = l_Lean_Syntax_node2(x_45, x_113, x_72, x_72); -x_115 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__43; +x_115 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__44; lean_inc(x_72); lean_inc(x_45); x_116 = l_Lean_Syntax_node4(x_45, x_115, x_112, x_57, x_114, x_72); if (lean_obj_tag(x_3) == 0) { lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; -x_117 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__48; +x_117 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__49; lean_inc(x_45); x_118 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_118, 0, x_45); @@ -4760,12 +4769,12 @@ x_120 = l_Lean_Syntax_node6(x_45, x_119, x_118, x_82, x_72, x_72, x_72, x_72); if (lean_obj_tag(x_42) == 0) { lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; -x_121 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__32; +x_121 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__33; x_122 = lean_array_push(x_121, x_63); -x_123 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__31; +x_123 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__32; x_124 = lean_array_push(x_122, x_123); x_125 = lean_box(2); -x_126 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__30; +x_126 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__31; x_127 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_127, 0, x_125); lean_ctor_set(x_127, 1, x_126); @@ -4785,12 +4794,12 @@ lean_dec(x_63); x_132 = lean_ctor_get(x_42, 0); lean_inc(x_132); lean_dec(x_42); -x_133 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__32; +x_133 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__33; x_134 = lean_array_push(x_133, x_132); -x_135 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__31; +x_135 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__32; x_136 = lean_array_push(x_134, x_135); x_137 = lean_box(2); -x_138 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__30; +x_138 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__31; x_139 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_139, 0, x_137); lean_ctor_set(x_139, 1, x_138); @@ -4824,12 +4833,12 @@ x_149 = l_Lean_Syntax_node6(x_45, x_148, x_147, x_82, x_72, x_72, x_72, x_72); if (lean_obj_tag(x_42) == 0) { lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; -x_150 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__32; +x_150 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__33; x_151 = lean_array_push(x_150, x_63); -x_152 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__31; +x_152 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__32; x_153 = lean_array_push(x_151, x_152); x_154 = lean_box(2); -x_155 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__30; +x_155 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__31; x_156 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_156, 0, x_154); lean_ctor_set(x_156, 1, x_155); @@ -4849,12 +4858,12 @@ lean_dec(x_63); x_161 = lean_ctor_get(x_42, 0); lean_inc(x_161); lean_dec(x_42); -x_162 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__32; +x_162 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__33; x_163 = lean_array_push(x_162, x_161); -x_164 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__31; +x_164 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__32; x_165 = lean_array_push(x_163, x_164); x_166 = lean_box(2); -x_167 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__30; +x_167 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__31; x_168 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_168, 0, x_166); lean_ctor_set(x_168, 1, x_167); @@ -4933,7 +4942,7 @@ lean_inc(x_45); x_198 = l_Lean_Syntax_node3(x_45, x_197, x_182, x_194, x_196); lean_inc(x_45); x_199 = l_Lean_Syntax_node1(x_45, x_187, x_198); -x_200 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__27; +x_200 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__29; lean_inc(x_45); x_201 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_201, 0, x_45); @@ -4941,12 +4950,12 @@ lean_ctor_set(x_201, 1, x_200); x_202 = lean_array_get_size(x_40); x_203 = lean_usize_of_nat(x_202); lean_dec(x_202); -x_204 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__35; +x_204 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__36; x_205 = l_Array_mapMUnsafe_map___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___spec__6(x_204, x_203, x_26, x_40); x_206 = lean_array_get_size(x_205); x_207 = lean_usize_of_nat(x_206); lean_dec(x_206); -x_208 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__36; +x_208 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__37; x_209 = l_Array_mapMUnsafe_map___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___spec__7(x_208, x_207, x_26, x_205); x_210 = l_Array_append___rarg(x_188, x_209); lean_inc(x_45); @@ -4959,7 +4968,7 @@ lean_inc(x_45); x_213 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_213, 0, x_45); lean_ctor_set(x_213, 1, x_212); -x_214 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__39; +x_214 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__40; lean_inc(x_45); x_215 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_215, 0, x_45); @@ -4969,12 +4978,12 @@ lean_inc(x_45); x_217 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_217, 0, x_45); lean_ctor_set(x_217, 1, x_216); -x_218 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__41; +x_218 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__42; lean_inc(x_45); x_219 = l_Lean_Syntax_node1(x_45, x_218, x_217); lean_inc(x_45); x_220 = l_Lean_Syntax_node1(x_45, x_187, x_219); -x_221 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__38; +x_221 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__39; lean_inc(x_45); x_222 = l_Lean_Syntax_node2(x_45, x_221, x_215, x_220); x_223 = l_Lean_expandExplicitBindersAux_loop___closed__11; @@ -4982,26 +4991,26 @@ lean_inc(x_45); x_224 = l_Lean_Syntax_node2(x_45, x_223, x_213, x_222); lean_inc(x_45); x_225 = l_Lean_Syntax_node1(x_45, x_187, x_224); -x_226 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__34; +x_226 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__35; lean_inc(x_45); x_227 = l_Lean_Syntax_node2(x_45, x_226, x_211, x_225); -x_228 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__44; +x_228 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__45; lean_inc(x_45); x_229 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_229, 0, x_45); lean_ctor_set(x_229, 1, x_228); -x_230 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__47; +x_230 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__48; lean_inc_n(x_189, 2); lean_inc(x_45); x_231 = l_Lean_Syntax_node2(x_45, x_230, x_189, x_189); -x_232 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__43; +x_232 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__44; lean_inc(x_189); lean_inc(x_45); x_233 = l_Lean_Syntax_node4(x_45, x_232, x_229, x_173, x_231, x_189); if (lean_obj_tag(x_3) == 0) { lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; -x_234 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__48; +x_234 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__49; lean_inc(x_45); x_235 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_235, 0, x_45); @@ -5014,12 +5023,12 @@ x_237 = l_Lean_Syntax_node6(x_45, x_236, x_235, x_199, x_189, x_189, x_189, x_18 if (lean_obj_tag(x_42) == 0) { lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; -x_238 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__32; +x_238 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__33; x_239 = lean_array_push(x_238, x_180); -x_240 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__31; +x_240 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__32; x_241 = lean_array_push(x_239, x_240); x_242 = lean_box(2); -x_243 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__30; +x_243 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__31; x_244 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_244, 0, x_242); lean_ctor_set(x_244, 1, x_243); @@ -5041,12 +5050,12 @@ lean_dec(x_180); x_250 = lean_ctor_get(x_42, 0); lean_inc(x_250); lean_dec(x_42); -x_251 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__32; +x_251 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__33; x_252 = lean_array_push(x_251, x_250); -x_253 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__31; +x_253 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__32; x_254 = lean_array_push(x_252, x_253); x_255 = lean_box(2); -x_256 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__30; +x_256 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__31; x_257 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_257, 0, x_255); lean_ctor_set(x_257, 1, x_256); @@ -5082,12 +5091,12 @@ x_268 = l_Lean_Syntax_node6(x_45, x_267, x_266, x_199, x_189, x_189, x_189, x_18 if (lean_obj_tag(x_42) == 0) { 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; -x_269 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__32; +x_269 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__33; x_270 = lean_array_push(x_269, x_180); -x_271 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__31; +x_271 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__32; x_272 = lean_array_push(x_270, x_271); x_273 = lean_box(2); -x_274 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__30; +x_274 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__31; x_275 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_275, 0, x_273); lean_ctor_set(x_275, 1, x_274); @@ -5109,12 +5118,12 @@ lean_dec(x_180); x_281 = lean_ctor_get(x_42, 0); lean_inc(x_281); lean_dec(x_42); -x_282 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__32; +x_282 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__33; x_283 = lean_array_push(x_282, x_281); -x_284 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__31; +x_284 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__32; x_285 = lean_array_push(x_283, x_284); x_286 = lean_box(2); -x_287 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__30; +x_287 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__31; x_288 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_288, 0, x_286); lean_ctor_set(x_288, 1, x_287); @@ -13669,7 +13678,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__45; lean_inc(x_33); x_46 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_46, 0, x_33); @@ -14004,7 +14013,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__29; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__30; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -14113,7 +14122,7 @@ 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__45; x_2 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; @@ -14650,7 +14659,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__48; +x_61 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__49; lean_inc(x_22); x_62 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_62, 0, x_22); @@ -14880,7 +14889,7 @@ 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___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__44; +x_29 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__45; lean_inc(x_19); x_30 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_30, 0, x_19); @@ -14893,7 +14902,7 @@ 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___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__48; +x_33 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__49; lean_inc(x_19); x_34 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_34, 0, x_19); @@ -15197,7 +15206,7 @@ lean_inc(x_42); x_55 = l_Lean_Syntax_node2(x_42, x_54, x_50, x_53); lean_inc(x_42); x_56 = l_Lean_Syntax_node1(x_42, x_47, x_55); -x_57 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__44; +x_57 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__45; lean_inc(x_42); x_58 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_58, 0, x_42); @@ -15323,7 +15332,7 @@ lean_inc(x_85); x_101 = l_Lean_Syntax_node2(x_85, x_73, x_100, x_77); lean_inc(x_85); x_102 = l_Lean_Syntax_node1(x_85, x_97, x_101); -x_103 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__44; +x_103 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__45; lean_inc(x_85); x_104 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_104, 0, x_85); @@ -15406,7 +15415,7 @@ lean_inc(x_118); x_138 = l_Lean_Syntax_node2(x_118, x_137, x_133, x_136); lean_inc(x_118); x_139 = l_Lean_Syntax_node1(x_118, x_130, x_138); -x_140 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__44; +x_140 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__45; lean_inc(x_118); x_141 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_141, 0, x_118); @@ -15565,7 +15574,7 @@ 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___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__44; +x_29 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__45; lean_inc(x_19); x_30 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_30, 0, x_19); @@ -15578,7 +15587,7 @@ 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___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__48; +x_33 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__49; lean_inc(x_19); x_34 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_34, 0, x_19); @@ -15797,7 +15806,7 @@ lean_inc(x_42); x_55 = l_Lean_Syntax_node2(x_42, x_54, x_50, x_53); lean_inc(x_42); x_56 = l_Lean_Syntax_node1(x_42, x_47, x_55); -x_57 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__44; +x_57 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__45; lean_inc(x_42); x_58 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_58, 0, x_42); @@ -15923,7 +15932,7 @@ lean_inc(x_85); x_101 = l_Lean_Syntax_node2(x_85, x_73, x_100, x_77); lean_inc(x_85); x_102 = l_Lean_Syntax_node1(x_85, x_97, x_101); -x_103 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__44; +x_103 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__45; lean_inc(x_85); x_104 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_104, 0, x_85); @@ -16006,7 +16015,7 @@ lean_inc(x_118); x_138 = l_Lean_Syntax_node2(x_118, x_137, x_133, x_136); lean_inc(x_118); x_139 = l_Lean_Syntax_node1(x_118, x_130, x_138); -x_140 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__44; +x_140 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__45; lean_inc(x_118); x_141 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_141, 0, x_118); @@ -19977,6 +19986,8 @@ l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint__ lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__47); l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__48 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__48(); lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__48); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__49 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__49(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c___x2d_u22a2____1___lambda__2___closed__49); 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(); diff --git a/stage0/stdlib/Init/Simproc.c b/stage0/stdlib/Init/Simproc.c index 29adc0aa5ae9..e403bda15536 100644 --- a/stage0/stdlib/Init/Simproc.c +++ b/stage0/stdlib/Init/Simproc.c @@ -293,6 +293,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Attr_sevalprocBuiltinAttr; static lean_object* l_Lean_Parser_command____Builtin__simproc_____x5b___x5d___x28___x29_x3a_x3d_____closed__2; static lean_object* l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__11; static lean_object* l_Lean_Parser_command____Simproc_____x5b___x5d___x28___x29_x3a_x3d_____closed__53; +static lean_object* l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__33; static lean_object* l_Array_forInUnsafe_loop___at___private_Init_Simproc_0__Lean_Parser_mkAttributeCmds___spec__1___closed__8; static lean_object* l_Lean_Parser_command____Simproc_____x5b___x5d___x28___x29_x3a_x3d_____closed__13; size_t lean_usize_add(size_t, size_t); @@ -2775,7 +2776,7 @@ static lean_object* _init_l_Lean_Parser___aux__Init__Simproc______macroRules__Le _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("def", 3); +x_1 = lean_mk_string_from_bytes("definition", 10); return x_1; } } @@ -2795,23 +2796,31 @@ static lean_object* _init_l_Lean_Parser___aux__Init__Simproc______macroRules__Le _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("declId", 6); +x_1 = lean_mk_string_from_bytes("def", 3); return x_1; } } static lean_object* _init_l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__16() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("declId", 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___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_command____Simproc_____x5b___x5d___x28___x29_x3a_x3d_____closed__1; x_2 = l_Lean_Parser_command____Simproc_____x5b___x5d___x28___x29_x3a_x3d_____closed__2; x_3 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__7; -x_4 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__15; +x_4 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___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___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__17() { +static lean_object* _init_l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__18() { _start: { lean_object* x_1; @@ -2819,19 +2828,19 @@ x_1 = lean_mk_string_from_bytes("optDeclSig", 10); return x_1; } } -static lean_object* _init_l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__18() { +static lean_object* _init_l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__19() { _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_command____Simproc_____x5b___x5d___x28___x29_x3a_x3d_____closed__1; x_2 = l_Lean_Parser_command____Simproc_____x5b___x5d___x28___x29_x3a_x3d_____closed__2; x_3 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__7; -x_4 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__17; +x_4 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__18; 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__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__19() { +static lean_object* _init_l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__20() { _start: { lean_object* x_1; @@ -2839,7 +2848,7 @@ x_1 = lean_mk_string_from_bytes("Term", 4); return x_1; } } -static lean_object* _init_l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__20() { +static lean_object* _init_l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__21() { _start: { lean_object* x_1; @@ -2847,19 +2856,19 @@ x_1 = lean_mk_string_from_bytes("typeSpec", 8); return x_1; } } -static lean_object* _init_l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__21() { +static lean_object* _init_l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__22() { _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_command____Simproc_____x5b___x5d___x28___x29_x3a_x3d_____closed__1; x_2 = l_Lean_Parser_command____Simproc_____x5b___x5d___x28___x29_x3a_x3d_____closed__2; -x_3 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__19; -x_4 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__20; +x_3 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__20; +x_4 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__21; 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__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__22() { +static lean_object* _init_l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__23() { _start: { lean_object* x_1; @@ -2867,7 +2876,7 @@ x_1 = lean_mk_string_from_bytes(":", 1); return x_1; } } -static lean_object* _init_l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__23() { +static lean_object* _init_l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__24() { _start: { lean_object* x_1; lean_object* x_2; @@ -2876,7 +2885,7 @@ x_2 = lean_mk_syntax_ident(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__24() { +static lean_object* _init_l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__25() { _start: { lean_object* x_1; @@ -2884,19 +2893,19 @@ x_1 = lean_mk_string_from_bytes("declValSimple", 13); return x_1; } } -static lean_object* _init_l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__25() { +static lean_object* _init_l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___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_Parser_command____Simproc_____x5b___x5d___x28___x29_x3a_x3d_____closed__1; x_2 = l_Lean_Parser_command____Simproc_____x5b___x5d___x28___x29_x3a_x3d_____closed__2; x_3 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__7; -x_4 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__24; +x_4 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___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_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__26() { +static lean_object* _init_l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__27() { _start: { lean_object* x_1; @@ -2904,7 +2913,7 @@ x_1 = lean_mk_string_from_bytes(":=", 2); return x_1; } } -static lean_object* _init_l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__27() { +static lean_object* _init_l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__28() { _start: { lean_object* x_1; @@ -2912,7 +2921,7 @@ x_1 = lean_mk_string_from_bytes("Termination", 11); return x_1; } } -static lean_object* _init_l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__28() { +static lean_object* _init_l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__29() { _start: { lean_object* x_1; @@ -2920,19 +2929,19 @@ x_1 = lean_mk_string_from_bytes("suffix", 6); return x_1; } } -static lean_object* _init_l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__29() { +static lean_object* _init_l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__30() { _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_command____Simproc_____x5b___x5d___x28___x29_x3a_x3d_____closed__1; x_2 = l_Lean_Parser_command____Simproc_____x5b___x5d___x28___x29_x3a_x3d_____closed__2; -x_3 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__27; -x_4 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__28; +x_3 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__28; +x_4 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__29; 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__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__30() { +static lean_object* _init_l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__31() { _start: { lean_object* x_1; @@ -2940,7 +2949,7 @@ x_1 = lean_mk_string_from_bytes("simproc_pattern%", 16); return x_1; } } -static lean_object* _init_l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__31() { +static lean_object* _init_l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__32() { _start: { lean_object* x_1; @@ -2948,7 +2957,7 @@ x_1 = lean_mk_string_from_bytes("=>", 2); return x_1; } } -static lean_object* _init_l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__32() { +static lean_object* _init_l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__33() { _start: { lean_object* x_1; lean_object* x_2; @@ -2997,41 +3006,41 @@ x_21 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_21, 0, x_18); lean_ctor_set(x_21, 1, x_19); lean_ctor_set(x_21, 2, x_20); -x_22 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__13; +x_22 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__15; lean_inc(x_18); x_23 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_23, 0, x_18); lean_ctor_set(x_23, 1, x_22); -x_24 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__16; +x_24 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__17; lean_inc(x_21); lean_inc(x_7); lean_inc(x_18); x_25 = l_Lean_Syntax_node2(x_18, x_24, x_7, x_21); -x_26 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__22; +x_26 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__23; lean_inc(x_18); x_27 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_27, 0, x_18); lean_ctor_set(x_27, 1, x_26); -x_28 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__21; -x_29 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__23; +x_28 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__22; +x_29 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__24; lean_inc(x_18); x_30 = l_Lean_Syntax_node2(x_18, x_28, x_27, x_29); lean_inc(x_18); x_31 = l_Lean_Syntax_node1(x_18, x_19, x_30); -x_32 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__18; +x_32 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__19; lean_inc(x_21); lean_inc(x_18); x_33 = l_Lean_Syntax_node2(x_18, x_32, x_21, x_31); -x_34 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__26; +x_34 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__27; lean_inc(x_18); x_35 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_35, 0, x_18); lean_ctor_set(x_35, 1, x_34); -x_36 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__29; +x_36 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__30; lean_inc_n(x_21, 2); lean_inc(x_18); x_37 = l_Lean_Syntax_node2(x_18, x_36, x_21, x_21); -x_38 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__25; +x_38 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__26; lean_inc(x_21); lean_inc(x_18); x_39 = l_Lean_Syntax_node4(x_18, x_38, x_35, x_15, x_37, x_21); @@ -3039,12 +3048,12 @@ x_40 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command lean_inc(x_21); lean_inc(x_18); x_41 = l_Lean_Syntax_node5(x_18, x_40, x_23, x_25, x_33, x_39, x_21); -x_42 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__30; +x_42 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__31; lean_inc(x_18); x_43 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_43, 0, x_18); lean_ctor_set(x_43, 1, x_42); -x_44 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__31; +x_44 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__32; lean_inc(x_18); x_45 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_45, 0, x_18); @@ -3055,7 +3064,7 @@ x_47 = l_Lean_Syntax_node4(x_18, x_46, x_43, x_13, x_45, x_7); if (lean_obj_tag(x_3) == 0) { lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_48 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__32; +x_48 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__33; lean_inc(x_18); x_49 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_49, 0, x_18); @@ -3281,41 +3290,41 @@ x_21 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_21, 0, x_18); lean_ctor_set(x_21, 1, x_19); lean_ctor_set(x_21, 2, x_20); -x_22 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__13; +x_22 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__15; lean_inc(x_18); x_23 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_23, 0, x_18); lean_ctor_set(x_23, 1, x_22); -x_24 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__16; +x_24 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__17; lean_inc(x_21); lean_inc(x_7); lean_inc(x_18); x_25 = l_Lean_Syntax_node2(x_18, x_24, x_7, x_21); -x_26 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__22; +x_26 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__23; lean_inc(x_18); x_27 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_27, 0, x_18); lean_ctor_set(x_27, 1, x_26); -x_28 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__21; +x_28 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__22; x_29 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Dsimproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__3; lean_inc(x_18); x_30 = l_Lean_Syntax_node2(x_18, x_28, x_27, x_29); lean_inc(x_18); x_31 = l_Lean_Syntax_node1(x_18, x_19, x_30); -x_32 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__18; +x_32 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__19; lean_inc(x_21); lean_inc(x_18); x_33 = l_Lean_Syntax_node2(x_18, x_32, x_21, x_31); -x_34 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__26; +x_34 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__27; lean_inc(x_18); x_35 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_35, 0, x_18); lean_ctor_set(x_35, 1, x_34); -x_36 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__29; +x_36 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__30; lean_inc_n(x_21, 2); lean_inc(x_18); x_37 = l_Lean_Syntax_node2(x_18, x_36, x_21, x_21); -x_38 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__25; +x_38 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__26; lean_inc(x_21); lean_inc(x_18); x_39 = l_Lean_Syntax_node4(x_18, x_38, x_35, x_15, x_37, x_21); @@ -3323,12 +3332,12 @@ x_40 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command lean_inc(x_21); lean_inc(x_18); x_41 = l_Lean_Syntax_node5(x_18, x_40, x_23, x_25, x_33, x_39, x_21); -x_42 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__30; +x_42 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__31; lean_inc(x_18); x_43 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_43, 0, x_18); lean_ctor_set(x_43, 1, x_42); -x_44 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__31; +x_44 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__32; lean_inc(x_18); x_45 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_45, 0, x_18); @@ -3339,7 +3348,7 @@ x_47 = l_Lean_Syntax_node4(x_18, x_46, x_43, x_13, x_45, x_7); if (lean_obj_tag(x_3) == 0) { lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_48 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__32; +x_48 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__33; lean_inc(x_18); x_49 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_49, 0, x_18); @@ -3532,41 +3541,41 @@ x_21 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_21, 0, x_18); lean_ctor_set(x_21, 1, x_19); lean_ctor_set(x_21, 2, x_20); -x_22 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__13; +x_22 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__15; lean_inc(x_18); x_23 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_23, 0, x_18); lean_ctor_set(x_23, 1, x_22); -x_24 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__16; +x_24 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__17; lean_inc(x_21); lean_inc(x_7); lean_inc(x_18); x_25 = l_Lean_Syntax_node2(x_18, x_24, x_7, x_21); -x_26 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__22; +x_26 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__23; lean_inc(x_18); x_27 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_27, 0, x_18); lean_ctor_set(x_27, 1, x_26); -x_28 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__21; -x_29 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__23; +x_28 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__22; +x_29 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__24; lean_inc(x_18); x_30 = l_Lean_Syntax_node2(x_18, x_28, x_27, x_29); lean_inc(x_18); x_31 = l_Lean_Syntax_node1(x_18, x_19, x_30); -x_32 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__18; +x_32 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__19; lean_inc(x_21); lean_inc(x_18); x_33 = l_Lean_Syntax_node2(x_18, x_32, x_21, x_31); -x_34 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__26; +x_34 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__27; lean_inc(x_18); x_35 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_35, 0, x_18); lean_ctor_set(x_35, 1, x_34); -x_36 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__29; +x_36 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__30; lean_inc_n(x_21, 2); lean_inc(x_18); x_37 = l_Lean_Syntax_node2(x_18, x_36, x_21, x_21); -x_38 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__25; +x_38 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__26; lean_inc(x_21); lean_inc(x_18); x_39 = l_Lean_Syntax_node4(x_18, x_38, x_35, x_15, x_37, x_21); @@ -3579,7 +3588,7 @@ lean_inc(x_18); x_43 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_43, 0, x_18); lean_ctor_set(x_43, 1, x_42); -x_44 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__31; +x_44 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__32; lean_inc(x_18); x_45 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_45, 0, x_18); @@ -3590,7 +3599,7 @@ x_47 = l_Lean_Syntax_node4(x_18, x_46, x_43, x_13, x_45, x_7); if (lean_obj_tag(x_3) == 0) { lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_48 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__32; +x_48 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__33; lean_inc(x_18); x_49 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_49, 0, x_18); @@ -3775,41 +3784,41 @@ x_21 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_21, 0, x_18); lean_ctor_set(x_21, 1, x_19); lean_ctor_set(x_21, 2, x_20); -x_22 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__13; +x_22 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__15; lean_inc(x_18); x_23 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_23, 0, x_18); lean_ctor_set(x_23, 1, x_22); -x_24 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__16; +x_24 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__17; lean_inc(x_21); lean_inc(x_7); lean_inc(x_18); x_25 = l_Lean_Syntax_node2(x_18, x_24, x_7, x_21); -x_26 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__22; +x_26 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__23; lean_inc(x_18); x_27 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_27, 0, x_18); lean_ctor_set(x_27, 1, x_26); -x_28 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__21; +x_28 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__22; x_29 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Dsimproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__3; lean_inc(x_18); x_30 = l_Lean_Syntax_node2(x_18, x_28, x_27, x_29); lean_inc(x_18); x_31 = l_Lean_Syntax_node1(x_18, x_19, x_30); -x_32 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__18; +x_32 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__19; lean_inc(x_21); lean_inc(x_18); x_33 = l_Lean_Syntax_node2(x_18, x_32, x_21, x_31); -x_34 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__26; +x_34 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__27; lean_inc(x_18); x_35 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_35, 0, x_18); lean_ctor_set(x_35, 1, x_34); -x_36 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__29; +x_36 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__30; lean_inc_n(x_21, 2); lean_inc(x_18); x_37 = l_Lean_Syntax_node2(x_18, x_36, x_21, x_21); -x_38 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__25; +x_38 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__26; lean_inc(x_21); lean_inc(x_18); x_39 = l_Lean_Syntax_node4(x_18, x_38, x_35, x_15, x_37, x_21); @@ -3822,7 +3831,7 @@ lean_inc(x_18); x_43 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_43, 0, x_18); lean_ctor_set(x_43, 1, x_42); -x_44 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__31; +x_44 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__32; lean_inc(x_18); x_45 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_45, 0, x_18); @@ -3833,7 +3842,7 @@ x_47 = l_Lean_Syntax_node4(x_18, x_46, x_43, x_13, x_45, x_7); if (lean_obj_tag(x_3) == 0) { lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_48 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__32; +x_48 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__33; lean_inc(x_18); x_49 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_49, 0, x_18); @@ -4021,7 +4030,7 @@ static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Init_Simproc_0 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_command____Simproc_____x5b___x5d___x28___x29_x3a_x3d_____closed__1; x_2 = l_Lean_Parser_command____Simproc_____x5b___x5d___x28___x29_x3a_x3d_____closed__2; -x_3 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__19; +x_3 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__20; x_4 = l_Array_forInUnsafe_loop___at___private_Init_Simproc_0__Lean_Parser_mkAttributeCmds___spec__1___closed__4; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -4298,7 +4307,7 @@ x_20 = l_Lean_Syntax_node1(x_10, x_19, x_4); if (lean_obj_tag(x_2) == 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; -x_21 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__32; +x_21 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__33; lean_inc(x_10); x_22 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_22, 0, x_10); @@ -4484,7 +4493,7 @@ lean_inc(x_21); x_27 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_27, 0, x_21); lean_ctor_set(x_27, 1, x_26); -x_28 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__26; +x_28 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__27; lean_inc(x_21); x_29 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_29, 0, x_21); @@ -4493,7 +4502,7 @@ if (lean_obj_tag(x_2) == 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_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__6; -x_31 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__32; +x_31 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__33; lean_inc(x_21); x_32 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_32, 0, x_21); @@ -4716,7 +4725,7 @@ static lean_object* _init_l_Lean_Parser___aux__Init__Simproc______macroRules__Le 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_command____Simproc_____x5b___x5d___x28___x29_x3a_x3d_____closed__1; x_2 = l_Lean_Parser_command____Simproc_____x5b___x5d___x28___x29_x3a_x3d_____closed__2; -x_3 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__19; +x_3 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__20; x_4 = l_Lean_Parser_command____Simproc_____x5b___x5d___x28___x29_x3a_x3d_____closed__13; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -4966,7 +4975,7 @@ lean_inc(x_21); x_27 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_27, 0, x_21); lean_ctor_set(x_27, 1, x_26); -x_28 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__26; +x_28 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__27; lean_inc(x_21); x_29 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_29, 0, x_21); @@ -4975,7 +4984,7 @@ if (lean_obj_tag(x_2) == 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_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__6; -x_31 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__32; +x_31 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__33; lean_inc(x_21); x_32 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_32, 0, x_21); @@ -5535,7 +5544,7 @@ lean_inc(x_45); x_51 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_51, 0, x_45); lean_ctor_set(x_51, 1, x_50); -x_52 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__26; +x_52 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__27; lean_inc(x_45); x_53 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_53, 0, x_45); @@ -5602,7 +5611,7 @@ x_71 = l_Lean_Syntax_node8(x_45, x_70, x_69, x_47, x_34, x_49, x_40, x_51, x_53, if (lean_obj_tag(x_5) == 0) { lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; -x_72 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__32; +x_72 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__33; lean_inc(x_45); x_73 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_73, 0, x_45); @@ -5761,7 +5770,7 @@ lean_inc(x_127); x_133 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_133, 0, x_127); lean_ctor_set(x_133, 1, x_132); -x_134 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__26; +x_134 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__27; lean_inc(x_127); x_135 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_135, 0, x_127); @@ -5823,7 +5832,7 @@ x_151 = l_Lean_Syntax_node8(x_127, x_150, x_149, x_129, x_116, x_131, x_122, x_1 if (lean_obj_tag(x_5) == 0) { lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; -x_152 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__32; +x_152 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__33; lean_inc(x_127); x_153 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_153, 0, x_127); @@ -5932,7 +5941,7 @@ lean_inc(x_192); x_198 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_198, 0, x_192); lean_ctor_set(x_198, 1, x_197); -x_199 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__26; +x_199 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__27; lean_inc(x_192); x_200 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_200, 0, x_192); @@ -5994,7 +6003,7 @@ x_216 = l_Lean_Syntax_node8(x_192, x_215, x_214, x_194, x_181, x_196, x_187, x_1 if (lean_obj_tag(x_5) == 0) { lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; -x_217 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__32; +x_217 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__33; lean_inc(x_192); x_218 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_218, 0, x_192); @@ -6385,7 +6394,7 @@ lean_inc(x_45); x_51 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_51, 0, x_45); lean_ctor_set(x_51, 1, x_50); -x_52 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__26; +x_52 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__27; lean_inc(x_45); x_53 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_53, 0, x_45); @@ -6452,7 +6461,7 @@ x_71 = l_Lean_Syntax_node8(x_45, x_70, x_69, x_47, x_34, x_49, x_40, x_51, x_53, if (lean_obj_tag(x_5) == 0) { lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; -x_72 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__32; +x_72 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__33; lean_inc(x_45); x_73 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_73, 0, x_45); @@ -6611,7 +6620,7 @@ lean_inc(x_127); x_133 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_133, 0, x_127); lean_ctor_set(x_133, 1, x_132); -x_134 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__26; +x_134 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__27; lean_inc(x_127); x_135 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_135, 0, x_127); @@ -6673,7 +6682,7 @@ x_151 = l_Lean_Syntax_node8(x_127, x_150, x_149, x_129, x_116, x_131, x_122, x_1 if (lean_obj_tag(x_5) == 0) { lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; -x_152 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__32; +x_152 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__33; lean_inc(x_127); x_153 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_153, 0, x_127); @@ -6782,7 +6791,7 @@ lean_inc(x_192); x_198 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_198, 0, x_192); lean_ctor_set(x_198, 1, x_197); -x_199 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__26; +x_199 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__27; lean_inc(x_192); x_200 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_200, 0, x_192); @@ -6844,7 +6853,7 @@ x_216 = l_Lean_Syntax_node8(x_192, x_215, x_214, x_194, x_181, x_196, x_187, x_1 if (lean_obj_tag(x_5) == 0) { lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; -x_217 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__32; +x_217 = l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__33; lean_inc(x_192); x_218 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_218, 0, x_192); @@ -7556,6 +7565,8 @@ l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simpr lean_mark_persistent(l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__31); l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__32 = _init_l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__32(); lean_mark_persistent(l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__32); +l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__33 = _init_l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__33(); +lean_mark_persistent(l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__33); l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___closed__1 = _init_l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___closed__1(); lean_mark_persistent(l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Simproc__decl___x28___x29_x3a_x3d____1___closed__1); l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Dsimproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__1 = _init_l_Lean_Parser___aux__Init__Simproc______macroRules__Lean__Parser__command__Dsimproc__decl___x28___x29_x3a_x3d____1___lambda__1___closed__1(); diff --git a/stage0/stdlib/Lean/Elab/Declaration.c b/stage0/stdlib/Lean/Elab/Declaration.c index be610870e98b..c227e8202f57 100644 --- a/stage0/stdlib/Lean/Elab/Declaration.c +++ b/stage0/stdlib/Lean/Elab/Declaration.c @@ -20,7 +20,6 @@ static lean_object* l_Lean_Elab_Command_expandInitialize___lambda__1___closed__7 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_Command_expandMutualNamespace(lean_object*); static lean_object* l_Lean_Elab_Command_expandInitialize___lambda__1___closed__20; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__11; static lean_object* l_Lean_Elab_Command_expandInitialize___lambda__1___closed__14; LEAN_EXPORT lean_object* l_Lean_Elab_Command_checkValidCtorModifier___at___private_Lean_Elab_Declaration_0__Lean_Elab_Command_inductiveSyntaxToView___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_setDefName___closed__3; @@ -43,7 +42,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Command_elabMutual_declRange___clos static lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_ensureValidNamespace___closed__3; lean_object* l_Lean_addDocString_x27___at_Lean_Elab_Command_expandDeclId___spec__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabDeclaration___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__12; static lean_object* l_Lean_Elab_Command_expandInitialize___lambda__1___closed__45; LEAN_EXPORT lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_expandDeclNamespace_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_expandMutualNamespace_declRange(lean_object*); @@ -82,7 +80,6 @@ LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_Command_expandMutualNamespace___ lean_object* l_Lean_Name_toString(lean_object*, uint8_t); LEAN_EXPORT lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_expandDeclNamespace_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getId(lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__7; static lean_object* l_Lean_Elab_Command_checkValidCtorModifier___at___private_Lean_Elab_Declaration_0__Lean_Elab_Command_inductiveSyntaxToView___spec__1___lambda__1___closed__2; static lean_object* l_Lean_Elab_Command_elabAxiom___lambda__5___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Command_expandMutualNamespace_declRange___closed__4; @@ -132,6 +129,7 @@ static lean_object* l_Lean_Elab_Command_elabDeclaration___closed__1; static lean_object* l_Lean_Elab_Command_expandInitialize___lambda__1___closed__12; static lean_object* l_Lean_Elab_Command_elabAxiom___lambda__5___closed__7; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Declaration_0__Lean_Elab_Command_inductiveSyntaxToView___spec__3___lambda__3(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_Declaration___hyg_7880____closed__5; LEAN_EXPORT lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_elabMutualInductive(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Command_elabAttr___spec__3___closed__2; static lean_object* l_Lean_Elab_Command_expandInitialize___lambda__1___closed__29; @@ -146,7 +144,6 @@ lean_object* l_Lean_Elab_elabModifiers___at_Lean_Elab_Command_elabMutualDef___sp static lean_object* l___regBuiltin_Lean_Elab_Command_elabDeclaration___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandInitialize___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isMutualPreambleCommand(lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__4; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabAttr___spec__4___lambda__1(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_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*); @@ -156,6 +153,8 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabDeclaration_declRange(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_expandMutualElement___closed__2; static lean_object* l_Lean_Elab_Command_expandInitialize___lambda__1___closed__3; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__4; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__11; 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___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__20; @@ -171,16 +170,17 @@ lean_object* l_Lean_CollectLevelParams_main(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__6; lean_object* l_Lean_Elab_Command_checkValidInductiveModifier___at_Lean_Elab_Command_elabStructure___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_expandDeclNamespace_x3f(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__23; lean_object* l_Lean_Elab_withSaveInfoContext___at___private_Lean_Elab_PreDefinition_Basic_0__Lean_Elab_addNonRecAux___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_checkValidCtorModifier___at___private_Lean_Elab_Declaration_0__Lean_Elab_Command_inductiveSyntaxToView___spec__1___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_expandInitialize_declRange(lean_object*); static lean_object* l_Lean_Elab_Command_elabDeclaration___closed__6; static lean_object* l___regBuiltin_Lean_Elab_Command_expandInitialize_declRange___closed__2; +static lean_object* l_Lean_Elab_Command_expandInitialize___lambda__1___closed__48; static lean_object* l_Lean_Elab_Command_checkValidCtorModifier___at___private_Lean_Elab_Declaration_0__Lean_Elab_Command_inductiveSyntaxToView___spec__1___lambda__2___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabAttr_declRange(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_expandInitialize(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabDeclaration_declRange___closed__1; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__8; static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabDeclaration___spec__4___closed__1; lean_object* l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSavedPartialInfoContext___at___private_Lean_Elab_PreDefinition_Basic_0__Lean_Elab_addNonRecAux___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_expandMutualPreamble___closed__3; @@ -205,6 +205,7 @@ uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_expandInitialize___closed__1; lean_object* l_Lean_Elab_expandMacroImpl_x3f(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_expandMutualNamespace___closed__2; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__2; lean_object* l_Lean_Elab_getDeclarationSelectionRef(lean_object*); lean_object* l_Lean_ResolveName_resolveNamespace(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__13; @@ -214,7 +215,7 @@ static lean_object* l_Lean_Elab_Command_elabAttr___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandInitialize___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_expandInitialize___lambda__1___closed__39; lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__13; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__9; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabDeclaration___spec__5(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_expandInitialize___lambda__1___closed__24; uint8_t l_Lean_Elab_Modifiers_isProtected(lean_object*); @@ -225,8 +226,6 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_expandMut LEAN_EXPORT uint8_t l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef(lean_object*); static lean_object* l_Lean_Elab_Command_expandInitialize___lambda__1___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Command_expandMutualNamespace_declRange___closed__2; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__2; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__6; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabAttr___spec__1___closed__4; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Declaration_0__Lean_Elab_Command_elabMutualInductive___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); 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*); @@ -252,6 +251,7 @@ static lean_object* l_Lean_Elab_Command_checkValidCtorModifier___at___private_Le lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_expandInitialize___lambda__1___closed__36; static lean_object* l_Lean_Elab_Command_elabAxiom___lambda__5___closed__3; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__13; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Declaration_0__Lean_Elab_Command_inductiveSyntaxToView___spec__3___closed__2; lean_object* l_Lean_Elab_elabAttrs___at_Lean_Elab_Command_elabMutualDef___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Command_liftTermElabM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -286,6 +286,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_getDeclarationRange___at_Lean_Elab_Command_ static lean_object* l_Lean_Elab_Command_expandInitialize___lambda__5___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_getDeclarationRange___at_Lean_Elab_Command_elabAxiom___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_expandInitialize___lambda__1___closed__34; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__7; static lean_object* l_Lean_Elab_Command_expandInitialize___lambda__2___closed__3; static lean_object* l_Lean_Elab_Command_expandInitialize___lambda__1___closed__26; static lean_object* l___regBuiltin_Lean_Elab_Command_expandMutualElement_declRange___closed__5; @@ -301,6 +302,8 @@ static lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_induc static lean_object* l___regBuiltin_Lean_Elab_Command_elabDeclaration_declRange___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Command_expandMutualPreamble_declRange___closed__6; lean_object* l_Lean_Elab_Command_getCurrMacroScope(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__1; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__12; static lean_object* l___regBuiltin_Lean_Elab_Command_expandMutualPreamble_declRange___closed__7; static lean_object* l_Lean_Elab_Command_expandInitialize___lambda__1___closed__16; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Declaration_0__Lean_Elab_Command_inductiveSyntaxToView___spec__3(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); @@ -337,7 +340,6 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabDeclaration(lean_o uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabDeclaration___spec__5___rarg___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandMutualNamespace___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__22; static lean_object* l___regBuiltin_Lean_Elab_Command_elabAttr_declRange___closed__1; uint8_t l_Lean_Syntax_isIdent(lean_object*); static lean_object* l_Lean_Elab_Command_elabMutual___closed__2; @@ -363,6 +365,7 @@ static lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isMut 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*); +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880_(lean_object*); lean_object* l_Lean_MessageData_ofExpr(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_expandMutualElement_declRange___closed__7; LEAN_EXPORT lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_splitMutualPreamble_loop(lean_object*, lean_object*); @@ -372,14 +375,12 @@ static lean_object* l___regBuiltin_Lean_Elab_Command_expandMutualElement___close static lean_object* l_Lean_Elab_Command_expandInitialize___lambda__1___closed__35; lean_object* l_panic___at_Lean_Parser_SyntaxStack_back___spec__1(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabDeclaration___closed__3; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__8; lean_object* l_Lean_Syntax_node4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabMutual___closed__2; static lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_setDeclIdName___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Command_elabAttr___closed__3; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Declaration_0__Lean_Elab_Command_elabMutualInductive___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__16; -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888_(lean_object*); lean_object* l_Lean_Syntax_setArg(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Elab_Modifiers_isPrivate(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Command_elabAttr___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -387,7 +388,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_ static lean_object* l___regBuiltin_Lean_Elab_Command_elabMutual_declRange___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Command_elabDeclaration_declRange___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Command_elabMutual_declRange___closed__2; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__3; 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_FileMap_leanPosToLspPos(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_expandInitialize___lambda__1___closed__23; @@ -401,7 +401,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_ static lean_object* l___regBuiltin_Lean_Elab_Command_elabAttr___closed__5; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Declaration_0__Lean_Elab_Command_inductiveSyntaxToView___spec__3___lambda__3___closed__2; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isMutualDef___spec__1(lean_object*, size_t, size_t); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__5; static lean_object* l_Lean_Elab_Command_checkValidCtorModifier___at___private_Lean_Elab_Declaration_0__Lean_Elab_Command_inductiveSyntaxToView___spec__1___lambda__2___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Command_expandMutualPreamble_declRange___closed__2; static lean_object* l_Lean_Elab_Command_expandInitialize___lambda__1___closed__38; @@ -413,6 +412,7 @@ static lean_object* l_Lean_Elab_Command_elabClassInductive___closed__1; static lean_object* l_Lean_Elab_Command_elabDeclaration___closed__8; LEAN_EXPORT lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_ensureValidNamespace(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabDeclaration(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__3; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isMutualDef___spec__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Command_elabAxiom___spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandMutualNamespace(lean_object*, lean_object*, lean_object*); @@ -433,7 +433,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_checkValidCtorModifier___at___priva static lean_object* l___regBuiltin_Lean_Elab_Command_expandMutualPreamble___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Command_expandInitialize_declRange___closed__1; lean_object* l_Lean_Elab_Command_elabInductiveViews(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__9; LEAN_EXPORT lean_object* l_Lean_addDeclarationRanges___at_Lean_Elab_Command_elabAxiom___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_setDeclIdName___closed__6; uint8_t l_Lean_isAttribute(lean_object*, lean_object*); @@ -451,7 +450,6 @@ static lean_object* l_Lean_Elab_Command_elabAxiom___lambda__4___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Command_expandInitialize_declRange___closed__5; static lean_object* l_Lean_Elab_Command_expandInitialize___lambda__1___closed__5; static lean_object* l_Lean_Elab_Command_expandMutualPreamble___closed__2; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__10; LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandMutualNamespace___lambda__1(lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabDeclaration___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__5; @@ -465,6 +463,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_checkValidCtorModifier___at___priva lean_object* l_Lean_Syntax_setInfo(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabAttr___spec__4(lean_object*, size_t, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ResolveName_resolveGlobalName(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__10; lean_object* l_Array_mkArray1___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandInitialize___lambda__5(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_expandInitialize___closed__3; @@ -491,7 +490,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Command_elabMutual_declRange___clos static lean_object* l_Lean_Elab_Command_elabDeclaration___closed__5; size_t lean_usize_add(size_t, size_t); static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabDeclaration___spec__5___rarg___closed__2; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__1; static lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandInitialize___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_Elab_Command_expandDeclId(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -523,6 +521,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_ static lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isMutualPreambleCommand___closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabAxiom___lambda__1___boxed(lean_object*); static lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_setDefName___closed__4; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandMutualElement(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Declaration_0__Lean_Elab_Command_inductiveSyntaxToView___spec__4(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_setDefName___closed__2; @@ -923,7 +922,7 @@ static lean_object* _init_l___private_Lean_Elab_Declaration_0__Lean_Elab_Command _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("def", 3); +x_1 = lean_mk_string_from_bytes("theorem", 7); return x_1; } } @@ -943,7 +942,7 @@ static lean_object* _init_l___private_Lean_Elab_Declaration_0__Lean_Elab_Command _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("theorem", 7); +x_1 = lean_mk_string_from_bytes("opaque", 6); return x_1; } } @@ -963,7 +962,7 @@ static lean_object* _init_l___private_Lean_Elab_Declaration_0__Lean_Elab_Command _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("opaque", 6); +x_1 = lean_mk_string_from_bytes("axiom", 5); return x_1; } } @@ -983,7 +982,7 @@ static lean_object* _init_l___private_Lean_Elab_Declaration_0__Lean_Elab_Command _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("axiom", 5); +x_1 = lean_mk_string_from_bytes("inductive", 9); return x_1; } } @@ -1003,7 +1002,7 @@ static lean_object* _init_l___private_Lean_Elab_Declaration_0__Lean_Elab_Command _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("inductive", 9); +x_1 = lean_mk_string_from_bytes("classInductive", 14); return x_1; } } @@ -1023,7 +1022,7 @@ static lean_object* _init_l___private_Lean_Elab_Declaration_0__Lean_Elab_Command _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("classInductive", 14); +x_1 = lean_mk_string_from_bytes("structure", 9); return x_1; } } @@ -1039,26 +1038,6 @@ x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__22() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("structure", 9); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___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___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__1; -x_2 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__2; -x_3 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__3; -x_4 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__22; -x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); -return x_5; -} -} LEAN_EXPORT uint8_t l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef(lean_object* x_1) { _start: { @@ -1117,14 +1096,25 @@ if (x_21 == 0) lean_object* x_22; uint8_t x_23; x_22 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__21; x_23 = lean_name_eq(x_7, x_22); -if (x_23 == 0) +lean_dec(x_7); +return x_23; +} +else { -lean_object* x_24; uint8_t x_25; -x_24 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__23; -x_25 = lean_name_eq(x_7, x_24); +uint8_t x_24; lean_dec(x_7); +x_24 = 1; +return x_24; +} +} +else +{ +uint8_t x_25; +lean_dec(x_7); +x_25 = 1; return x_25; } +} else { uint8_t x_26; @@ -1165,30 +1155,6 @@ x_30 = 1; return x_30; } } -else -{ -uint8_t x_31; -lean_dec(x_7); -x_31 = 1; -return x_31; -} -} -else -{ -uint8_t x_32; -lean_dec(x_7); -x_32 = 1; -return x_32; -} -} -else -{ -uint8_t x_33; -lean_dec(x_7); -x_33 = 1; -return x_33; -} -} } } LEAN_EXPORT lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___boxed(lean_object* x_1) { @@ -1360,7 +1326,7 @@ static lean_object* _init_l___private_Lean_Elab_Declaration_0__Lean_Elab_Command 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_Declaration_0__Lean_Elab_Command_setDeclIdName___closed__5; x_2 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_setDefName___closed__3; -x_3 = lean_unsigned_to_nat(82u); +x_3 = lean_unsigned_to_nat(81u); x_4 = lean_unsigned_to_nat(4u); x_5 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_setDefName___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -2279,7 +2245,7 @@ static lean_object* _init_l_Lean_Elab_Command_elabAxiom___lambda__5___closed__6( { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Command_elabAxiom___lambda__5___closed__5; -x_2 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__16; +x_2 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__14; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -5212,22 +5178,22 @@ x_9 = lean_unsigned_to_nat(1u); x_10 = l_Lean_Syntax_getArg(x_1, x_9); lean_inc(x_10); x_11 = l_Lean_Syntax_getKind(x_10); -x_12 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__17; +x_12 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__15; x_13 = lean_name_eq(x_11, x_12); if (x_13 == 0) { lean_object* x_14; uint8_t x_15; -x_14 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__19; +x_14 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__17; x_15 = lean_name_eq(x_11, x_14); if (x_15 == 0) { lean_object* x_16; uint8_t x_17; -x_16 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__21; +x_16 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__19; x_17 = lean_name_eq(x_11, x_16); if (x_17 == 0) { lean_object* x_18; uint8_t x_19; -x_18 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__23; +x_18 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__21; x_19 = lean_name_eq(x_11, x_18); lean_dec(x_11); if (x_19 == 0) @@ -5647,7 +5613,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabDeclaration_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(193u); +x_1 = lean_unsigned_to_nat(192u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -5659,7 +5625,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabDeclaration_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(217u); +x_1 = lean_unsigned_to_nat(216u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -5687,7 +5653,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabDeclaration_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(193u); +x_1 = lean_unsigned_to_nat(192u); x_2 = lean_unsigned_to_nat(4u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -5699,7 +5665,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabDeclaration_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(193u); +x_1 = lean_unsigned_to_nat(192u); x_2 = lean_unsigned_to_nat(19u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -5758,7 +5724,7 @@ x_6 = lean_unsigned_to_nat(1u); x_7 = l_Lean_Syntax_getArg(x_5, x_6); lean_dec(x_5); x_8 = l_Lean_Syntax_getKind(x_7); -x_9 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__19; +x_9 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__17; x_10 = lean_name_eq(x_8, x_9); lean_dec(x_8); if (x_10 == 0) @@ -6643,7 +6609,7 @@ static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_expandMu 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_Declaration_0__Lean_Elab_Command_setDeclIdName___closed__5; x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_expandMutualNamespace___spec__3___closed__1; -x_3 = lean_unsigned_to_nat(295u); +x_3 = lean_unsigned_to_nat(294u); x_4 = lean_unsigned_to_nat(40u); x_5 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_expandMutualNamespace___spec__3___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -7124,7 +7090,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandMutualNamespace _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(286u); +x_1 = lean_unsigned_to_nat(285u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7136,7 +7102,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandMutualNamespace _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(301u); +x_1 = lean_unsigned_to_nat(300u); x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7164,7 +7130,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandMutualNamespace _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(286u); +x_1 = lean_unsigned_to_nat(285u); x_2 = lean_unsigned_to_nat(4u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7176,7 +7142,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandMutualNamespace _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(286u); +x_1 = lean_unsigned_to_nat(285u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7602,7 +7568,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandMutualElement_d _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(304u); +x_1 = lean_unsigned_to_nat(303u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7614,7 +7580,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandMutualElement_d _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(314u); +x_1 = lean_unsigned_to_nat(313u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7642,7 +7608,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandMutualElement_d _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(304u); +x_1 = lean_unsigned_to_nat(303u); x_2 = lean_unsigned_to_nat(4u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7654,7 +7620,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandMutualElement_d _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(304u); +x_1 = lean_unsigned_to_nat(303u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7848,7 +7814,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandMutualPreamble_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(317u); +x_1 = lean_unsigned_to_nat(316u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7860,7 +7826,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandMutualPreamble_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(324u); +x_1 = lean_unsigned_to_nat(323u); x_2 = lean_unsigned_to_nat(74u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7888,7 +7854,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandMutualPreamble_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(317u); +x_1 = lean_unsigned_to_nat(316u); x_2 = lean_unsigned_to_nat(4u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7900,7 +7866,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandMutualPreamble_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(317u); +x_1 = lean_unsigned_to_nat(316u); x_2 = lean_unsigned_to_nat(24u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -8056,7 +8022,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabMutual_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(327u); +x_1 = lean_unsigned_to_nat(326u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -8068,7 +8034,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabMutual_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(333u); +x_1 = lean_unsigned_to_nat(332u); x_2 = lean_unsigned_to_nat(152u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -8096,7 +8062,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabMutual_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(327u); +x_1 = lean_unsigned_to_nat(326u); x_2 = lean_unsigned_to_nat(4u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -8108,7 +8074,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabMutual_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(327u); +x_1 = lean_unsigned_to_nat(326u); x_2 = lean_unsigned_to_nat(14u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9392,7 +9358,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabAttr_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(336u); +x_1 = lean_unsigned_to_nat(335u); x_2 = lean_unsigned_to_nat(36u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9404,7 +9370,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabAttr_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(368u); +x_1 = lean_unsigned_to_nat(367u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9432,7 +9398,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabAttr_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(336u); +x_1 = lean_unsigned_to_nat(335u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9444,7 +9410,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabAttr_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(336u); +x_1 = lean_unsigned_to_nat(335u); x_2 = lean_unsigned_to_nat(48u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9606,23 +9572,31 @@ static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___clo _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("declId", 6); +x_1 = lean_mk_string_from_bytes("def", 3); return x_1; } } static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__14() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("declId", 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___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___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__1; x_2 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__2; x_3 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__3; -x_4 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__13; +x_4 = l_Lean_Elab_Command_expandInitialize___lambda__1___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_Elab_Command_expandInitialize___lambda__1___closed__15() { +static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__16() { _start: { lean_object* x_1; @@ -9630,26 +9604,26 @@ x_1 = lean_mk_string_from_bytes("initFn", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__16() { +static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__17() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__15; +x_1 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__16; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__17() { +static lean_object* _init_l_Lean_Elab_Command_expandInitialize___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_Elab_Command_expandInitialize___lambda__1___closed__15; +x_2 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__16; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__18() { +static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__19() { _start: { lean_object* x_1; @@ -9657,19 +9631,19 @@ x_1 = lean_mk_string_from_bytes("optDeclSig", 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__19() { +static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__1; x_2 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__2; x_3 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__3; -x_4 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__18; +x_4 = l_Lean_Elab_Command_expandInitialize___lambda__1___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_Elab_Command_expandInitialize___lambda__1___closed__20() { +static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__21() { _start: { lean_object* x_1; @@ -9677,19 +9651,19 @@ x_1 = lean_mk_string_from_bytes("typeSpec", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__21() { +static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__1; x_2 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__2; x_3 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__2; -x_4 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__20; +x_4 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__21; 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_expandInitialize___lambda__1___closed__22() { +static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__23() { _start: { lean_object* x_1; @@ -9697,7 +9671,7 @@ x_1 = lean_mk_string_from_bytes(":", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__23() { +static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__24() { _start: { lean_object* x_1; @@ -9705,19 +9679,19 @@ x_1 = lean_mk_string_from_bytes("app", 3); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__24() { +static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__1; x_2 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__2; x_3 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__2; -x_4 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__23; +x_4 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__24; 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_expandInitialize___lambda__1___closed__25() { +static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__26() { _start: { lean_object* x_1; @@ -9725,72 +9699,72 @@ x_1 = lean_mk_string_from_bytes("IO", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__26() { +static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__27() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__25; +x_1 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__26; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__27() { +static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__28() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__25; +x_2 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__26; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__28() { +static lean_object* _init_l_Lean_Elab_Command_expandInitialize___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_Elab_Command_expandInitialize___lambda__1___closed__27; +x_2 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__28; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__29() { +static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__30() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__27; +x_1 = l_Lean_Elab_Command_expandInitialize___lambda__1___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_Elab_Command_expandInitialize___lambda__1___closed__30() { +static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__31() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__29; +x_2 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__30; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__31() { +static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__32() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__28; -x_2 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__30; +x_1 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__29; +x_2 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__31; 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_Elab_Command_expandInitialize___lambda__1___closed__32() { +static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__33() { _start: { lean_object* x_1; @@ -9798,72 +9772,72 @@ x_1 = lean_mk_string_from_bytes("Unit", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__33() { +static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__34() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__32; +x_1 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__33; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__34() { +static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__35() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__32; +x_2 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__33; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__35() { +static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__36() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__34; +x_2 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__35; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__36() { +static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__37() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__34; +x_1 = l_Lean_Elab_Command_expandInitialize___lambda__1___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_Elab_Command_expandInitialize___lambda__1___closed__37() { +static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__38() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__36; +x_2 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__37; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__38() { +static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__39() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__35; -x_2 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__37; +x_1 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__36; +x_2 = l_Lean_Elab_Command_expandInitialize___lambda__1___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_Elab_Command_expandInitialize___lambda__1___closed__39() { +static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__40() { _start: { lean_object* x_1; @@ -9871,19 +9845,19 @@ x_1 = lean_mk_string_from_bytes("declValSimple", 13); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__40() { +static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___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___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__1; x_2 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__2; x_3 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__3; -x_4 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__39; +x_4 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__40; 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_expandInitialize___lambda__1___closed__41() { +static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__42() { _start: { lean_object* x_1; @@ -9891,7 +9865,7 @@ x_1 = lean_mk_string_from_bytes(":=", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__42() { +static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__43() { _start: { lean_object* x_1; @@ -9899,19 +9873,19 @@ x_1 = lean_mk_string_from_bytes("do", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__43() { +static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___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___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__1; x_2 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__2; x_3 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__2; -x_4 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__42; +x_4 = l_Lean_Elab_Command_expandInitialize___lambda__1___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_Command_expandInitialize___lambda__1___closed__44() { +static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__45() { _start: { lean_object* x_1; @@ -9919,7 +9893,7 @@ x_1 = lean_mk_string_from_bytes("Termination", 11); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__45() { +static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__46() { _start: { lean_object* x_1; @@ -9927,19 +9901,19 @@ x_1 = lean_mk_string_from_bytes("suffix", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__46() { +static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___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___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__1; x_2 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__2; -x_3 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__44; -x_4 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__45; +x_3 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__45; +x_4 = l_Lean_Elab_Command_expandInitialize___lambda__1___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_Command_expandInitialize___lambda__1___closed__47() { +static lean_object* _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__48() { _start: { lean_object* x_1; lean_object* x_2; @@ -10082,48 +10056,48 @@ lean_inc(x_37); x_55 = l_Lean_Syntax_node3(x_37, x_54, x_41, x_51, x_53); lean_inc(x_37); x_56 = l_Lean_Syntax_node1(x_37, x_42, x_55); -x_57 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__10; +x_57 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__13; lean_inc(x_37); x_58 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_58, 0, x_37); lean_ctor_set(x_58, 1, x_57); -x_59 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__17; +x_59 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__18; lean_inc(x_38); lean_inc(x_39); x_60 = l_Lean_addMacroScope(x_39, x_59, x_38); x_61 = lean_box(0); -x_62 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__16; +x_62 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__17; lean_inc(x_37); x_63 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_63, 0, x_37); lean_ctor_set(x_63, 1, x_62); lean_ctor_set(x_63, 2, x_60); lean_ctor_set(x_63, 3, x_61); -x_64 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__14; +x_64 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__15; lean_inc(x_44); lean_inc(x_37); x_65 = l_Lean_Syntax_node2(x_37, x_64, x_63, x_44); -x_66 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__22; +x_66 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__23; lean_inc(x_37); x_67 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_67, 0, x_37); lean_ctor_set(x_67, 1, x_66); -x_68 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__27; +x_68 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__28; lean_inc(x_38); lean_inc(x_39); x_69 = l_Lean_addMacroScope(x_39, x_68, x_38); -x_70 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__26; -x_71 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__31; +x_70 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__27; +x_71 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__32; lean_inc(x_37); x_72 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_72, 0, x_37); lean_ctor_set(x_72, 1, x_70); lean_ctor_set(x_72, 2, x_69); lean_ctor_set(x_72, 3, x_71); -x_73 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__34; +x_73 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__35; x_74 = l_Lean_addMacroScope(x_39, x_73, x_38); -x_75 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__33; -x_76 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__38; +x_75 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__34; +x_76 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__39; lean_inc(x_37); x_77 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_77, 0, x_37); @@ -10132,36 +10106,36 @@ lean_ctor_set(x_77, 2, x_74); lean_ctor_set(x_77, 3, x_76); lean_inc(x_37); x_78 = l_Lean_Syntax_node1(x_37, x_42, x_77); -x_79 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__24; +x_79 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__25; lean_inc(x_37); x_80 = l_Lean_Syntax_node2(x_37, x_79, x_72, x_78); -x_81 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__21; +x_81 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__22; lean_inc(x_37); x_82 = l_Lean_Syntax_node2(x_37, x_81, x_67, x_80); lean_inc(x_37); x_83 = l_Lean_Syntax_node1(x_37, x_42, x_82); -x_84 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__19; +x_84 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__20; lean_inc(x_44); lean_inc(x_37); x_85 = l_Lean_Syntax_node2(x_37, x_84, x_44, x_83); -x_86 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__41; +x_86 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__42; lean_inc(x_37); x_87 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_87, 0, x_37); lean_ctor_set(x_87, 1, x_86); -x_88 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__42; +x_88 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__43; lean_inc(x_37); x_89 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_89, 0, x_37); lean_ctor_set(x_89, 1, x_88); -x_90 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__43; +x_90 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__44; lean_inc(x_37); x_91 = l_Lean_Syntax_node2(x_37, x_90, x_89, x_3); -x_92 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__46; +x_92 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__47; lean_inc_n(x_44, 2); lean_inc(x_37); x_93 = l_Lean_Syntax_node2(x_37, x_92, x_44, x_44); -x_94 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__40; +x_94 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__41; lean_inc(x_44); lean_inc(x_37); x_95 = l_Lean_Syntax_node4(x_37, x_94, x_87, x_91, x_93, x_44); @@ -10172,7 +10146,7 @@ x_97 = l_Lean_Syntax_node5(x_37, x_96, x_58, x_65, x_85, x_95, x_44); if (lean_obj_tag(x_6) == 0) { 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_98 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__47; +x_98 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__48; lean_inc(x_37); x_99 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_99, 0, x_37); @@ -10400,37 +10374,37 @@ x_22 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_22, 0, x_17); lean_ctor_set(x_22, 1, x_20); lean_ctor_set(x_22, 2, x_21); -x_23 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__10; +x_23 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__13; lean_inc(x_17); x_24 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_24, 0, x_17); lean_ctor_set(x_24, 1, x_23); -x_25 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__17; +x_25 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__18; lean_inc(x_18); lean_inc(x_19); x_26 = l_Lean_addMacroScope(x_19, x_25, x_18); x_27 = lean_box(0); -x_28 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__16; +x_28 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__17; lean_inc(x_17); x_29 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_29, 0, x_17); lean_ctor_set(x_29, 1, x_28); lean_ctor_set(x_29, 2, x_26); lean_ctor_set(x_29, 3, x_27); -x_30 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__14; +x_30 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__15; lean_inc(x_22); lean_inc(x_29); lean_inc(x_17); x_31 = l_Lean_Syntax_node2(x_17, x_30, x_29, x_22); -x_32 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__22; +x_32 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__23; lean_inc(x_17); x_33 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_33, 0, x_17); lean_ctor_set(x_33, 1, x_32); -x_34 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__27; +x_34 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__28; x_35 = l_Lean_addMacroScope(x_19, x_34, x_18); -x_36 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__26; -x_37 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__31; +x_36 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__27; +x_37 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__32; lean_inc(x_17); x_38 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_38, 0, x_17); @@ -10440,20 +10414,20 @@ lean_ctor_set(x_38, 3, x_37); lean_inc(x_1); lean_inc(x_17); x_39 = l_Lean_Syntax_node1(x_17, x_20, x_1); -x_40 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__24; +x_40 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__25; lean_inc(x_17); x_41 = l_Lean_Syntax_node2(x_17, x_40, x_38, x_39); -x_42 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__21; +x_42 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__22; lean_inc(x_33); lean_inc(x_17); x_43 = l_Lean_Syntax_node2(x_17, x_42, x_33, x_41); lean_inc(x_17); x_44 = l_Lean_Syntax_node1(x_17, x_20, x_43); -x_45 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__19; +x_45 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__20; lean_inc(x_22); lean_inc(x_17); x_46 = l_Lean_Syntax_node2(x_17, x_45, x_22, x_44); -x_47 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__41; +x_47 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__42; lean_inc(x_17); x_48 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_48, 0, x_17); @@ -10470,23 +10444,23 @@ lean_ctor_set(x_52, 0, x_17); lean_ctor_set(x_52, 1, x_51); lean_inc(x_17); x_53 = l_Lean_Syntax_node1(x_17, x_20, x_52); -x_54 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__42; +x_54 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__43; lean_inc(x_17); x_55 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_55, 0, x_17); lean_ctor_set(x_55, 1, x_54); -x_56 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__43; +x_56 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__44; lean_inc(x_17); x_57 = l_Lean_Syntax_node2(x_17, x_56, x_55, x_2); x_58 = l_Lean_Elab_Command_expandInitialize___lambda__2___closed__2; lean_inc(x_3); lean_inc(x_17); x_59 = l_Lean_Syntax_node4(x_17, x_58, x_50, x_53, x_3, x_57); -x_60 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__46; +x_60 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__47; lean_inc_n(x_22, 2); lean_inc(x_17); x_61 = l_Lean_Syntax_node2(x_17, x_60, x_22, x_22); -x_62 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__40; +x_62 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__41; lean_inc(x_22); lean_inc(x_17); x_63 = l_Lean_Syntax_node4(x_17, x_62, x_48, x_59, x_61, x_22); @@ -10522,7 +10496,7 @@ lean_inc(x_17); x_79 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_79, 0, x_17); lean_ctor_set(x_79, 1, x_78); -x_80 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__14; +x_80 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__12; lean_inc(x_17); x_81 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_81, 0, x_17); @@ -10542,7 +10516,7 @@ x_89 = l_Lean_Elab_Command_expandInitialize___lambda__2___closed__9; lean_inc(x_22); lean_inc(x_17); x_90 = l_Lean_Syntax_node2(x_17, x_89, x_22, x_88); -x_91 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__15; +x_91 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__13; lean_inc(x_22); lean_inc(x_17); x_92 = l_Lean_Syntax_node4(x_17, x_91, x_81, x_87, x_90, x_22); @@ -10641,7 +10615,7 @@ x_107 = l_Lean_Syntax_node1(x_17, x_20, x_106); if (lean_obj_tag(x_14) == 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; -x_108 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__47; +x_108 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__48; lean_inc(x_17); x_109 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_109, 0, x_17); @@ -11485,7 +11459,7 @@ lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; x_27 = l_Lean_Syntax_getArg(x_21, x_8); x_28 = l_Lean_Syntax_getArg(x_21, x_14); lean_dec(x_21); -x_29 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__21; +x_29 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__22; lean_inc(x_28); x_30 = l_Lean_Syntax_isOfKind(x_28, x_29); if (x_30 == 0) @@ -11597,7 +11571,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandInitialize_decl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(370u); +x_1 = lean_unsigned_to_nat(369u); x_2 = lean_unsigned_to_nat(50u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11609,7 +11583,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandInitialize_decl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(382u); +x_1 = lean_unsigned_to_nat(381u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11637,7 +11611,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandInitialize_decl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(370u); +x_1 = lean_unsigned_to_nat(369u); x_2 = lean_unsigned_to_nat(54u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11649,7 +11623,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandInitialize_decl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(370u); +x_1 = lean_unsigned_to_nat(369u); x_2 = lean_unsigned_to_nat(70u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11695,7 +11669,7 @@ x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__1() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -11705,37 +11679,37 @@ 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_Declaration___hyg_7888____closed__2() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____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_Declaration___hyg_7888____closed__1; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__1; x_2 = l_Lean_Elab_Command_elabAxiom___lambda__5___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_Declaration___hyg_7888____closed__3() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____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_Declaration___hyg_7888____closed__2; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__2; x_2 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__3; 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_Declaration___hyg_7888____closed__4() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____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_Declaration___hyg_7888____closed__3; -x_2 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__15; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__3; +x_2 = l_Lean_Elab_Command_expandInitialize___lambda__1___closed__16; 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_Declaration___hyg_7888____closed__5() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__5() { _start: { lean_object* x_1; @@ -11743,37 +11717,37 @@ x_1 = lean_mk_string_from_bytes("_@", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__6() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____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_Declaration___hyg_7888____closed__4; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__5; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__4; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____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_Declaration___hyg_7888____closed__7() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____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_Declaration___hyg_7888____closed__6; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__6; x_2 = l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___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_Declaration___hyg_7888____closed__8() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____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_Declaration___hyg_7888____closed__7; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__7; x_2 = l_Lean_Elab_Command_elabAxiom___lambda__5___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_Declaration___hyg_7888____closed__9() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__9() { _start: { lean_object* x_1; @@ -11781,17 +11755,17 @@ x_1 = lean_mk_string_from_bytes("Declaration", 11); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__10() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____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_Declaration___hyg_7888____closed__8; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__9; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__8; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____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_Declaration___hyg_7888____closed__11() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__11() { _start: { lean_object* x_1; @@ -11799,33 +11773,33 @@ x_1 = lean_mk_string_from_bytes("_hyg", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__12() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____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_Declaration___hyg_7888____closed__10; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__11; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__10; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__11; 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_Declaration___hyg_7888____closed__13() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____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_Declaration___hyg_7888____closed__12; -x_2 = lean_unsigned_to_nat(7888u); +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__12; +x_2 = lean_unsigned_to_nat(7880u); 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_Declaration___hyg_7888_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880_(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_elabAxiom___lambda__5___closed__6; x_3 = 0; -x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__13; +x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__13; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -11927,10 +11901,6 @@ l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__20 = lean_mark_persistent(l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__20); l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__21 = _init_l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__21(); lean_mark_persistent(l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__21); -l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__22 = _init_l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__22(); -lean_mark_persistent(l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__22); -l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__23 = _init_l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__23(); -lean_mark_persistent(l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__23); l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isInstanceDef___closed__1 = _init_l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isInstanceDef___closed__1(); lean_mark_persistent(l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isInstanceDef___closed__1); l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isInstanceDef___closed__2 = _init_l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isInstanceDef___closed__2(); @@ -12362,6 +12332,8 @@ l_Lean_Elab_Command_expandInitialize___lambda__1___closed__46 = _init_l_Lean_Ela lean_mark_persistent(l_Lean_Elab_Command_expandInitialize___lambda__1___closed__46); l_Lean_Elab_Command_expandInitialize___lambda__1___closed__47 = _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__47(); lean_mark_persistent(l_Lean_Elab_Command_expandInitialize___lambda__1___closed__47); +l_Lean_Elab_Command_expandInitialize___lambda__1___closed__48 = _init_l_Lean_Elab_Command_expandInitialize___lambda__1___closed__48(); +lean_mark_persistent(l_Lean_Elab_Command_expandInitialize___lambda__1___closed__48); l_Lean_Elab_Command_expandInitialize___lambda__2___closed__1 = _init_l_Lean_Elab_Command_expandInitialize___lambda__2___closed__1(); lean_mark_persistent(l_Lean_Elab_Command_expandInitialize___lambda__2___closed__1); l_Lean_Elab_Command_expandInitialize___lambda__2___closed__2 = _init_l_Lean_Elab_Command_expandInitialize___lambda__2___closed__2(); @@ -12434,33 +12406,33 @@ lean_mark_persistent(l___regBuiltin_Lean_Elab_Command_expandInitialize_declRange if (builtin) {res = l___regBuiltin_Lean_Elab_Command_expandInitialize_declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__1(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__1); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__2(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__2); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__3(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__3); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__4 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__4(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__4); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__5 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__5(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__5); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__6 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__6(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__6); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__7 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__7(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__7); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__8 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__8(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__8); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__9 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__9(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__9); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__10 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__10(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__10); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__11 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__11(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__11); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__12 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__12(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__12); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__13 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__13(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888____closed__13); -if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7888_(lean_io_mk_world()); +}l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__1(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__1); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__2(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__2); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__3(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__3); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__4 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__4(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__4); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__5 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__5(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__5); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__6 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__6(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__6); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__7 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__7(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__7); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__8 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__8(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__8); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__9 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__9(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__9); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__10 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__10(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__10); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__11 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__11(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__11); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__12 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__12(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__12); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__13 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__13(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880____closed__13); +if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Declaration___hyg_7880_(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/DefView.c b/stage0/stdlib/Lean/Elab/DefView.c index b109160054b9..5c744cc6ae8b 100644 --- a/stage0/stdlib/Lean/Elab/DefView.c +++ b/stage0/stdlib/Lean/Elab/DefView.c @@ -14,9 +14,9 @@ extern "C" { #endif LEAN_EXPORT lean_object* l_Lean_Meta_visitLambda_visit___at_Lean_Elab_Command_mkInstanceName___spec__5___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_mkInstanceName___spec__14___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withoutErrToSorry___at_Lean_Elab_Command_mkInstanceName___spec__13(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_DefView___hyg_2224____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_mkInstanceName___spec__14___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_mkInstanceName___spec__14___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint32_t lean_string_utf8_get(lean_object*, lean_object*); @@ -31,7 +31,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkDefViewOfExample(lean_object*, le LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkDefViewOfInstance(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forEachExpr_x27_visit___at_Lean_Elab_Command_mkInstanceName___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* l_Lean_Elab_Command_runTermElabM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__9; LEAN_EXPORT lean_object* l_Lean_Meta_forEachExpr_x27_visit___at_Lean_Elab_Command_mkInstanceName___spec__3___lambda__1(lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_maxRecDepthErrorMessage; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_mkInstanceName___spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -44,7 +43,6 @@ static lean_object* l_Lean_Elab_Command_mkInstanceName___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkDefViewOfOpaque___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkInstanceName___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_Meta_forEachExpr_x27___at_Lean_Elab_Command_mkInstanceName___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_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__11; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_mkDefViewOfInstance___spec__2(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_Elab_liftMacroM___at_Lean_Elab_Command_mkInstanceName___spec__14(lean_object*, lean_object*, lean_object*, lean_object*); @@ -55,12 +53,12 @@ static lean_object* l_Lean_Elab_Command_isDefLike___closed__5; uint8_t lean_usize_dec_eq(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Meta_visitForall_visit___at_Lean_Elab_Command_mkInstanceName___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*); static lean_object* l_Lean_Elab_Command_mkDefViewOfInstance___closed__6; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__11; static lean_object* l_Lean_Elab_Command_mkDefViewOfInstance___closed__10; lean_object* l_Lean_Syntax_getArgs(lean_object*); lean_object* l_ST_Prim_Ref_get___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkDefViewOfInstance___lambda__1___boxed(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_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_mkDefViewOfInstance___spec__5___boxed(lean_object*, lean_object*); 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_Meta_withLocalDecl___at_Lean_Elab_Command_mkInstanceName___spec__6___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*); @@ -77,11 +75,11 @@ lean_object* l_Lean_stringToMessageData(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_mkDefViewOfInstance___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_isDefLike___closed__9; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_mkInstanceName___spec__14___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__10; uint8_t lean_string_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_mkDefViewOfOpaque___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkDefViewOfOpaque___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_mkDefViewOfOpaque___lambda__2___closed__2; -static lean_object* l_Lean_Elab_Command_isDefLike___closed__12; LEAN_EXPORT uint8_t l_Lean_Elab_DefKind_isExample(uint8_t); LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkDefViewOfOpaque___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_visitLet_visit___at_Lean_Elab_Command_mkInstanceName___spec__11___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -103,7 +101,6 @@ static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_mkInstanceN LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_mkDefViewOfInstance___spec__5(lean_object*, lean_object*); uint8_t l_Lean_Expr_isSort(lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__7; 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_Lean_Elab_liftMacroM___at_Lean_Elab_Command_mkInstanceName___spec__14___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_liftExcept___at_Lean_Elab_liftMacroM___spec__1(lean_object*, lean_object*, lean_object*); @@ -117,13 +114,14 @@ lean_object* l_Lean_ResolveName_resolveNamespace(lean_object*, lean_object*, lea LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_DefView_isInstance___spec__1(lean_object*, size_t, size_t); static lean_object* l_Lean_Elab_Command_mkDefViewOfAbbrev___closed__1; lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__15; LEAN_EXPORT uint8_t l_Lean_Elab_DefKind_isTheorem(uint8_t); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_mkInstanceName___spec__14___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_instInhabitedDefView___closed__2; lean_object* l_Lean_Elab_Term_withoutErrToSorryImp___rarg(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_DefView___hyg_2224____closed__8; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_mkDefView___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_mkInstanceName___spec__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__5; lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_mkDefViewOfInstance___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_mkInstanceName___spec__16(lean_object*, lean_object*, lean_object*, lean_object*); @@ -137,9 +135,9 @@ LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_mkInstanc lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_visitLambda_visit___at_Lean_Elab_Command_mkInstanceName___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_mkInstanceName___spec__15(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__13; LEAN_EXPORT lean_object* l_Lean_Elab_DefKind_toCtorIdx___boxed(lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__16; LEAN_EXPORT uint8_t l_Lean_Elab_Command_isDefLike(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_mkDefViewOfInstance___spec__8(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getOptional_x3f(lean_object*); @@ -153,6 +151,7 @@ lean_object* l_Lean_Expr_constName_x21(lean_object*); lean_object* l_Lean_Elab_Command_getCurrMacroScope(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_HashMapImp_find_x3f___at_Lean_ForEachExpr_visit___spec__1(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Elab_DefView_isInstance(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224_(lean_object*); static lean_object* l_Lean_Elab_Command_mkDefViewOfExample___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_DefKind_toCtorIdx(uint8_t); static lean_object* l_Lean_Elab_Command_mkDefViewOfAbbrev___closed__4; @@ -169,19 +168,24 @@ LEAN_EXPORT lean_object* l_Lean_Elab_DefKind_isExample___boxed(lean_object*); 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*); LEAN_EXPORT lean_object* l_Lean_Meta_visitLet___at_Lean_Elab_Command_mkInstanceName___spec__10(lean_object*, lean_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_isDefLike___closed__13; LEAN_EXPORT lean_object* l_Lean_Elab_instInhabitedDefView; LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkInstanceName___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__9; lean_object* l_Lean_Elab_Command_getScope___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__3; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__4; static lean_object* l_Lean_Meta_forEachExpr_x27___at_Lean_Elab_Command_mkInstanceName___spec__2___closed__1; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__7; static lean_object* l_Lean_Elab_Command_mkDefViewOfInstance___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_forEachExpr___at_Lean_Elab_Command_mkInstanceName___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__6; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__12; uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_visitForall_visit___at_Lean_Elab_Command_mkInstanceName___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* l_Lean_Elab_expandOptNamedPrio___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_visitLet_visit___at_Lean_Elab_Command_mkInstanceName___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* 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_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__14; LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkDefView(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkInstanceName___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_isDefLike___boxed(lean_object*); @@ -205,7 +209,6 @@ uint8_t lean_nat_dec_lt(lean_object*, lean_object*); lean_object* lean_environment_main_module(lean_object*); static lean_object* l_Lean_Elab_instInhabitedDefView___closed__3; static lean_object* l_Lean_Elab_Command_mkDefViewOfOpaque___lambda__2___closed__4; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__10; 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_mkDefViewOfInstance___closed__8; @@ -213,10 +216,10 @@ 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*); static lean_object* l_Lean_Elab_Command_isDefLike___closed__8; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__15; lean_object* l_Lean_Syntax_node1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_mkDefViewOfExample___closed__2; LEAN_EXPORT uint8_t l___private_Lean_Elab_DefView_0__Lean_Elab_beqDefKind____x40_Lean_Elab_DefView___hyg_15_(uint8_t, uint8_t); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__2; uint8_t l_Lean_Syntax_isNone(lean_object*); static lean_object* l_Lean_Elab_Command_mkDefViewOfOpaque___lambda__2___closed__5; static lean_object* l_Lean_Elab_Command_isDefLike___closed__6; @@ -227,10 +230,9 @@ static lean_object* l_Lean_Elab_Command_mkDefViewOfInstance___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_mkDefViewOfInstance___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_mkInstanceName___lambda__2___closed__1; uint8_t l_Lean_Expr_isProp(lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__12; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_Command_mkInstanceName___spec__6(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__6; lean_object* l_Lean_Elab_expandOptDeclSig(lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__13; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_mkDefViewOfInstance___spec__10___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Elab_Command_mkInstanceName___spec__12___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_erase_macro_scopes(lean_object*); @@ -247,8 +249,6 @@ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_DefView_isInstanc lean_object* l_String_capitalize(lean_object*); static lean_object* l_Lean_Elab_Command_isDefLike___closed__4; size_t lean_usize_add(size_t, size_t); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__3; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__4; static lean_object* l_Lean_Elab_Command_mkDefViewOfOpaque___lambda__2___closed__6; static lean_object* l_Lean_Elab_Command_mkDefViewOfOpaque___closed__2; lean_object* lean_array_uget(lean_object*, size_t); @@ -257,12 +257,11 @@ lean_object* l_Lean_Elab_mkUnusedBaseName(lean_object*, lean_object*, lean_objec 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*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__14; lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_mkInstanceName___lambda__1___closed__2; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__1; static lean_object* l_Lean_Elab_Command_mkDefViewOfInstance___closed__11; LEAN_EXPORT lean_object* l_Lean_Elab_DefKind_noConfusion___rarg___lambda__1(lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__16; LEAN_EXPORT lean_object* l___private_Lean_Elab_DefView_0__Lean_Elab_beqDefKind____x40_Lean_Elab_DefView___hyg_15____boxed(lean_object*, lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_DefKind_isTheorem___boxed(lean_object*); @@ -295,7 +294,6 @@ lean_object* l_Lean_Elab_Term_elabType(lean_object*, lean_object*, lean_object*, lean_object* l_Lean_Syntax_mkNumLit(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_Meta_withLocalDecl___at_Lean_Elab_Command_mkInstanceName___spec__6___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*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkDefViewOfDef(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_DefKind_toCtorIdx(uint8_t x_1) { _start: @@ -6107,7 +6105,7 @@ static lean_object* _init_l_Lean_Elab_Command_isDefLike___closed__5() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("def", 3); +x_1 = lean_mk_string_from_bytes("theorem", 7); return x_1; } } @@ -6127,7 +6125,7 @@ static lean_object* _init_l_Lean_Elab_Command_isDefLike___closed__7() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("theorem", 7); +x_1 = lean_mk_string_from_bytes("opaque", 6); return x_1; } } @@ -6146,26 +6144,6 @@ return x_5; static lean_object* _init_l_Lean_Elab_Command_isDefLike___closed__9() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("opaque", 6); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_Command_isDefLike___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_Elab_Command_mkDefViewOfInstance___closed__1; -x_2 = l_Lean_Elab_Command_mkDefViewOfInstance___closed__2; -x_3 = l_Lean_Elab_Command_mkDefViewOfInstance___closed__7; -x_4 = l_Lean_Elab_Command_isDefLike___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_Elab_Command_isDefLike___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_Elab_Command_mkDefViewOfInstance___closed__1; x_2 = l_Lean_Elab_Command_mkDefViewOfInstance___closed__2; @@ -6175,7 +6153,7 @@ 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_isDefLike___closed__12() { +static lean_object* _init_l_Lean_Elab_Command_isDefLike___closed__10() { _start: { lean_object* x_1; @@ -6183,14 +6161,14 @@ x_1 = lean_mk_string_from_bytes("example", 7); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_isDefLike___closed__13() { +static lean_object* _init_l_Lean_Elab_Command_isDefLike___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_Elab_Command_mkDefViewOfInstance___closed__1; x_2 = l_Lean_Elab_Command_mkDefViewOfInstance___closed__2; x_3 = l_Lean_Elab_Command_mkDefViewOfInstance___closed__7; -x_4 = l_Lean_Elab_Command_isDefLike___closed__12; +x_4 = l_Lean_Elab_Command_isDefLike___closed__10; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } @@ -6220,21 +6198,32 @@ x_10 = lean_name_eq(x_2, x_9); if (x_10 == 0) { lean_object* x_11; uint8_t x_12; -x_11 = l_Lean_Elab_Command_isDefLike___closed__10; +x_11 = l_Lean_Elab_Command_isDefLike___closed__9; x_12 = lean_name_eq(x_2, x_11); if (x_12 == 0) { lean_object* x_13; uint8_t x_14; x_13 = l_Lean_Elab_Command_isDefLike___closed__11; x_14 = lean_name_eq(x_2, x_13); -if (x_14 == 0) +lean_dec(x_2); +return x_14; +} +else { -lean_object* x_15; uint8_t x_16; -x_15 = l_Lean_Elab_Command_isDefLike___closed__13; -x_16 = lean_name_eq(x_2, x_15); +uint8_t x_15; +lean_dec(x_2); +x_15 = 1; +return x_15; +} +} +else +{ +uint8_t x_16; lean_dec(x_2); +x_16 = 1; return x_16; } +} else { uint8_t x_17; @@ -6259,30 +6248,6 @@ x_19 = 1; return x_19; } } -else -{ -uint8_t x_20; -lean_dec(x_2); -x_20 = 1; -return x_20; -} -} -else -{ -uint8_t x_21; -lean_dec(x_2); -x_21 = 1; -return x_21; -} -} -else -{ -uint8_t x_22; -lean_dec(x_2); -x_22 = 1; -return x_22; -} -} } LEAN_EXPORT lean_object* l_Lean_Elab_Command_isDefLike___boxed(lean_object* x_1) { _start: @@ -6389,54 +6354,62 @@ x_14 = lean_name_eq(x_6, x_13); if (x_14 == 0) { lean_object* x_15; uint8_t x_16; -x_15 = l_Lean_Elab_Command_isDefLike___closed__10; +x_15 = l_Lean_Elab_Command_isDefLike___closed__9; x_16 = lean_name_eq(x_6, x_15); if (x_16 == 0) { lean_object* x_17; uint8_t x_18; x_17 = l_Lean_Elab_Command_isDefLike___closed__11; x_18 = lean_name_eq(x_6, x_17); -if (x_18 == 0) -{ -lean_object* x_19; uint8_t x_20; -x_19 = l_Lean_Elab_Command_isDefLike___closed__13; -x_20 = lean_name_eq(x_6, x_19); lean_dec(x_6); -if (x_20 == 0) +if (x_18 == 0) { -lean_object* x_21; lean_object* x_22; +lean_object* x_19; lean_object* x_20; lean_dec(x_2); lean_dec(x_1); -x_21 = l_Lean_Elab_Command_mkDefView___closed__2; -x_22 = l_Lean_throwError___at_Lean_Elab_Command_mkDefView___spec__1(x_21, x_3, x_4, x_5); +x_19 = l_Lean_Elab_Command_mkDefView___closed__2; +x_20 = l_Lean_throwError___at_Lean_Elab_Command_mkDefView___spec__1(x_19, x_3, x_4, x_5); lean_dec(x_4); -return x_22; +return x_20; } else { -lean_object* x_23; lean_object* x_24; +lean_object* x_21; lean_object* x_22; lean_dec(x_4); lean_dec(x_3); -x_23 = l_Lean_Elab_Command_mkDefViewOfExample(x_1, x_2); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_5); -return x_24; +x_21 = l_Lean_Elab_Command_mkDefViewOfExample(x_1, x_2); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_5); +return x_22; } } else { -lean_object* x_25; +lean_object* x_23; lean_dec(x_6); -x_25 = l_Lean_Elab_Command_mkDefViewOfInstance(x_1, x_2, x_3, x_4, x_5); -return x_25; +x_23 = l_Lean_Elab_Command_mkDefViewOfInstance(x_1, x_2, x_3, x_4, x_5); +return x_23; } } else { -lean_object* x_26; +lean_object* x_24; lean_dec(x_6); -x_26 = l_Lean_Elab_Command_mkDefViewOfOpaque(x_1, x_2, x_3, x_4, x_5); +x_24 = l_Lean_Elab_Command_mkDefViewOfOpaque(x_1, x_2, x_3, x_4, x_5); +return x_24; +} +} +else +{ +lean_object* x_25; lean_object* x_26; +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +x_25 = l_Lean_Elab_Command_mkDefViewOfTheorem(x_1, x_2); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_5); return x_26; } } @@ -6446,7 +6419,7 @@ lean_object* x_27; lean_object* x_28; lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); -x_27 = l_Lean_Elab_Command_mkDefViewOfTheorem(x_1, x_2); +x_27 = l_Lean_Elab_Command_mkDefViewOfDef(x_1, x_2); x_28 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_28, 0, x_27); lean_ctor_set(x_28, 1, x_5); @@ -6459,39 +6432,13 @@ lean_object* x_29; lean_object* x_30; lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); -x_29 = l_Lean_Elab_Command_mkDefViewOfDef(x_1, x_2); +x_29 = l_Lean_Elab_Command_mkDefViewOfAbbrev(x_1, x_2); x_30 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_30, 0, x_29); lean_ctor_set(x_30, 1, x_5); return x_30; } } -else -{ -lean_object* x_31; lean_object* x_32; -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_3); -x_31 = l_Lean_Elab_Command_mkDefViewOfDef(x_1, x_2); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_5); -return x_32; -} -} -else -{ -lean_object* x_33; lean_object* x_34; -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_3); -x_33 = l_Lean_Elab_Command_mkDefViewOfAbbrev(x_1, x_2); -x_34 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_34, 0, x_33); -lean_ctor_set(x_34, 1, x_5); -return x_34; -} -} } LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_mkDefView___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: @@ -6502,7 +6449,7 @@ lean_dec(x_3); return x_5; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__1() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__1() { _start: { lean_object* x_1; @@ -6510,17 +6457,17 @@ x_1 = lean_mk_string_from_bytes("Elab", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__2() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____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_DefView___hyg_2255____closed__1; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__1; x_2 = l_Lean_Elab_Command_isDefLike___closed__3; 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_DefView___hyg_2255____closed__3() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -6530,27 +6477,27 @@ 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_DefView___hyg_2255____closed__4() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__3; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__1; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__3; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____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_DefView___hyg_2255____closed__5() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____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_DefView___hyg_2255____closed__4; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__4; x_2 = l_Lean_Elab_Command_mkDefViewOfInstance___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_DefView___hyg_2255____closed__6() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__6() { _start: { lean_object* x_1; @@ -6558,17 +6505,17 @@ x_1 = lean_mk_string_from_bytes("initFn", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__7() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____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_DefView___hyg_2255____closed__5; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__6; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__5; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____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_DefView___hyg_2255____closed__8() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__8() { _start: { lean_object* x_1; @@ -6576,37 +6523,37 @@ x_1 = lean_mk_string_from_bytes("_@", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__9() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____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_DefView___hyg_2255____closed__7; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__8; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__7; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____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_DefView___hyg_2255____closed__10() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____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_DefView___hyg_2255____closed__9; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__9; x_2 = l_Lean_Elab_Command_mkDefViewOfInstance___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_DefView___hyg_2255____closed__11() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____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_DefView___hyg_2255____closed__10; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__1; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__10; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____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_DefView___hyg_2255____closed__12() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__12() { _start: { lean_object* x_1; @@ -6614,17 +6561,17 @@ x_1 = lean_mk_string_from_bytes("DefView", 7); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__13() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____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_DefView___hyg_2255____closed__11; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__12; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__11; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____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_DefView___hyg_2255____closed__14() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__14() { _start: { lean_object* x_1; @@ -6632,33 +6579,33 @@ x_1 = lean_mk_string_from_bytes("_hyg", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__15() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__13; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__14; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__13; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__14; 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_DefView___hyg_2255____closed__16() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____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_DefView___hyg_2255____closed__15; -x_2 = lean_unsigned_to_nat(2255u); +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__15; +x_2 = lean_unsigned_to_nat(2224u); 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_DefView___hyg_2255_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224_(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_DefView___hyg_2255____closed__2; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__2; x_3 = 0; -x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__16; +x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__16; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -6809,47 +6756,43 @@ l_Lean_Elab_Command_isDefLike___closed__10 = _init_l_Lean_Elab_Command_isDefLike lean_mark_persistent(l_Lean_Elab_Command_isDefLike___closed__10); l_Lean_Elab_Command_isDefLike___closed__11 = _init_l_Lean_Elab_Command_isDefLike___closed__11(); lean_mark_persistent(l_Lean_Elab_Command_isDefLike___closed__11); -l_Lean_Elab_Command_isDefLike___closed__12 = _init_l_Lean_Elab_Command_isDefLike___closed__12(); -lean_mark_persistent(l_Lean_Elab_Command_isDefLike___closed__12); -l_Lean_Elab_Command_isDefLike___closed__13 = _init_l_Lean_Elab_Command_isDefLike___closed__13(); -lean_mark_persistent(l_Lean_Elab_Command_isDefLike___closed__13); l_Lean_Elab_Command_mkDefView___closed__1 = _init_l_Lean_Elab_Command_mkDefView___closed__1(); lean_mark_persistent(l_Lean_Elab_Command_mkDefView___closed__1); l_Lean_Elab_Command_mkDefView___closed__2 = _init_l_Lean_Elab_Command_mkDefView___closed__2(); lean_mark_persistent(l_Lean_Elab_Command_mkDefView___closed__2); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__1(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__1); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__2(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__2); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__3(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__3); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__4 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__4(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__4); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__5 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__5(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__5); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__6 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__6(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__6); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__7 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__7(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__7); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__8 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__8(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__8); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__9 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__9(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__9); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__10 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__10(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__10); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__11 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__11(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__11); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__12 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__12(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__12); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__13 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__13(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__13); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__14 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__14(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__14); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__15 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__15(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__15); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__16 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__16(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255____closed__16); -if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2255_(lean_io_mk_world()); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__1(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__1); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__2(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__2); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__3(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__3); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__4 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__4(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__4); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__5 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__5(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__5); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__6 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__6(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__6); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__7 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__7(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__7); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__8 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__8(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__8); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__9 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__9(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__9); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__10 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__10(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__10); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__11 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__11(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__11); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__12 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__12(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__12); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__13 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__13(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__13); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__14 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__14(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__14); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__15 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__15(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__15); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__16 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__16(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224____closed__16); +if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2224_(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.c b/stage0/stdlib/Lean/Elab/Quotation.c index 88ea69714418..2883377ed905 100644 --- a/stage0/stdlib/Lean/Elab/Quotation.c +++ b/stage0/stdlib/Lean/Elab/Quotation.c @@ -132,6 +132,7 @@ static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_ static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__7; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__14; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__38; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__60; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabNoErrorIfUnused(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t lean_uint64_of_nat(lean_object*); @@ -16955,7 +16956,7 @@ static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotatio _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("def", 3); +x_1 = lean_mk_string_from_bytes("definition", 10); return x_1; } } @@ -16975,23 +16976,31 @@ static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotatio _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("declId", 6); +x_1 = lean_mk_string_from_bytes("def", 3); return x_1; } } static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__22() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("declId", 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__1; x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__2; x_3 = l_Lean_Elab_Term_Quotation_getQuotKind___closed__5; -x_4 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__21; +x_4 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____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_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__23() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__24() { _start: { lean_object* x_1; @@ -16999,26 +17008,26 @@ x_1 = lean_mk_string_from_bytes("elabQuot", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__24() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__25() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__23; +x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__24; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__25() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__26() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__23; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__24; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__26() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__27() { _start: { lean_object* x_1; @@ -17026,19 +17035,19 @@ x_1 = lean_mk_string_from_bytes("optDeclSig", 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__27() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__28() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__1; x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__2; x_3 = l_Lean_Elab_Term_Quotation_getQuotKind___closed__5; -x_4 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__26; +x_4 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___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_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__28() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__29() { _start: { lean_object* x_1; @@ -17046,19 +17055,19 @@ x_1 = lean_mk_string_from_bytes("typeSpec", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__29() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__30() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__1; x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__2; x_3 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__3; -x_4 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__28; +x_4 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__29; 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_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__30() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__31() { _start: { lean_object* x_1; @@ -17066,7 +17075,7 @@ x_1 = lean_mk_string_from_bytes(":", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__31() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__32() { _start: { lean_object* x_1; @@ -17074,62 +17083,62 @@ x_1 = lean_mk_string_from_bytes("TermElab", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__32() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__33() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__31; +x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__32; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__33() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__34() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__31; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__32; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__34() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__35() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__1; x_2 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__1; x_3 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__3; -x_4 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__31; +x_4 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__32; 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_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__35() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__36() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__34; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__35; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__36() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__37() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__35; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__36; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__37() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__38() { _start: { lean_object* x_1; @@ -17137,19 +17146,19 @@ x_1 = lean_mk_string_from_bytes("declValSimple", 13); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__38() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__39() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__1; x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__2; x_3 = l_Lean_Elab_Term_Quotation_getQuotKind___closed__5; -x_4 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__37; +x_4 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__38; 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_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__39() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__40() { _start: { lean_object* x_1; @@ -17157,84 +17166,84 @@ x_1 = lean_mk_string_from_bytes("adaptExpander", 13); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__40() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__41() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__39; +x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__40; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__41() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__42() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__39; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__40; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__42() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__43() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__1; x_2 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__1; x_3 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__3; -x_4 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__39; +x_4 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__40; 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_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__43() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__44() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__42; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__43; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__44() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__45() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__42; +x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__43; 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_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__45() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__46() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__44; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__45; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__46() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__47() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__43; -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__45; +x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__44; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__46; 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_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__47() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__48() { _start: { lean_object* x_1; @@ -17242,16 +17251,16 @@ x_1 = lean_mk_string_from_bytes("stxQuot.expand", 14); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__48() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__49() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__47; +x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__48; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__49() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__50() { _start: { lean_object* x_1; @@ -17259,7 +17268,7 @@ x_1 = lean_mk_string_from_bytes("stxQuot", 7); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__50() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__51() { _start: { lean_object* x_1; @@ -17267,17 +17276,17 @@ x_1 = lean_mk_string_from_bytes("expand", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__51() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__52() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__49; -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__50; +x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__50; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__51; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__52() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__53() { _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; @@ -17285,59 +17294,59 @@ x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiqu x_2 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__1; x_3 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__3; x_4 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__2; -x_5 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__49; -x_6 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__50; +x_5 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__50; +x_6 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__51; x_7 = l_Lean_Name_mkStr6(x_1, x_2, x_3, x_4, x_5, x_6); return x_7; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__53() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__54() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__52; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__53; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__54() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__55() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__52; +x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__53; 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_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__55() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__56() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__54; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__55; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__56() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__57() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__53; -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__55; +x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__54; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__56; 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_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__57() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__58() { _start: { lean_object* x_1; @@ -17345,7 +17354,7 @@ x_1 = lean_mk_string_from_bytes("Termination", 11); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__58() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__59() { _start: { lean_object* x_1; @@ -17353,14 +17362,14 @@ x_1 = lean_mk_string_from_bytes("suffix", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__59() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__60() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__1; x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__2; -x_3 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__57; -x_4 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__58; +x_3 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__58; +x_4 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__59; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } @@ -17450,49 +17459,49 @@ x_38 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__ lean_inc_n(x_17, 5); lean_inc(x_12); x_39 = l_Lean_Syntax_node6(x_12, x_38, x_17, x_37, x_17, x_17, x_17, x_17); -x_40 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__19; +x_40 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__21; lean_inc(x_12); x_41 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_41, 0, x_12); lean_ctor_set(x_41, 1, x_40); -x_42 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__25; +x_42 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__26; lean_inc(x_13); lean_inc(x_14); x_43 = l_Lean_addMacroScope(x_14, x_42, x_13); -x_44 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__24; +x_44 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__25; lean_inc(x_12); x_45 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_45, 0, x_12); lean_ctor_set(x_45, 1, x_44); lean_ctor_set(x_45, 2, x_43); lean_ctor_set(x_45, 3, x_24); -x_46 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__22; +x_46 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__23; lean_inc(x_17); lean_inc(x_12); x_47 = l_Lean_Syntax_node2(x_12, x_46, x_45, x_17); -x_48 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__30; +x_48 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__31; lean_inc(x_12); x_49 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_49, 0, x_12); lean_ctor_set(x_49, 1, x_48); -x_50 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__33; +x_50 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__34; lean_inc(x_13); lean_inc(x_14); x_51 = l_Lean_addMacroScope(x_14, x_50, x_13); -x_52 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__32; -x_53 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__36; +x_52 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__33; +x_53 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__37; lean_inc(x_12); x_54 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_54, 0, x_12); lean_ctor_set(x_54, 1, x_52); lean_ctor_set(x_54, 2, x_51); lean_ctor_set(x_54, 3, x_53); -x_55 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__29; +x_55 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__30; lean_inc(x_12); x_56 = l_Lean_Syntax_node2(x_12, x_55, x_49, x_54); lean_inc(x_12); x_57 = l_Lean_Syntax_node1(x_12, x_15, x_56); -x_58 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__27; +x_58 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__28; lean_inc(x_17); lean_inc(x_12); x_59 = l_Lean_Syntax_node2(x_12, x_58, x_17, x_57); @@ -17501,22 +17510,22 @@ lean_inc(x_12); x_61 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_61, 0, x_12); lean_ctor_set(x_61, 1, x_60); -x_62 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__41; +x_62 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__42; lean_inc(x_13); lean_inc(x_14); x_63 = l_Lean_addMacroScope(x_14, x_62, x_13); -x_64 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__40; -x_65 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__46; +x_64 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__41; +x_65 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__47; lean_inc(x_12); x_66 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_66, 0, x_12); lean_ctor_set(x_66, 1, x_64); lean_ctor_set(x_66, 2, x_63); lean_ctor_set(x_66, 3, x_65); -x_67 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__51; +x_67 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__52; x_68 = l_Lean_addMacroScope(x_14, x_67, x_13); -x_69 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__48; -x_70 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__56; +x_69 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__49; +x_70 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__57; lean_inc(x_12); x_71 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_71, 0, x_12); @@ -17528,11 +17537,11 @@ x_72 = l_Lean_Syntax_node1(x_12, x_15, x_71); x_73 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__6; lean_inc(x_12); x_74 = l_Lean_Syntax_node2(x_12, x_73, x_66, x_72); -x_75 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__59; +x_75 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__60; lean_inc_n(x_17, 2); lean_inc(x_12); x_76 = l_Lean_Syntax_node2(x_12, x_75, x_17, x_17); -x_77 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__38; +x_77 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__39; lean_inc(x_17); lean_inc(x_12); x_78 = l_Lean_Syntax_node4(x_12, x_77, x_61, x_74, x_76, x_17); @@ -17622,7 +17631,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x4 { 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_6357____closed__5; -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__23; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__24; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -39080,6 +39089,8 @@ l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__E lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__58); l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__59 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__59(); lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__59); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__60 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__60(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__60); l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_6357____closed__1 = _init_l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_6357____closed__1(); lean_mark_persistent(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_6357____closed__1); l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_6357____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_6357____closed__1(); diff --git a/stage0/stdlib/Lean/Elab/Tactic/Config.c b/stage0/stdlib/Lean/Elab/Tactic/Config.c index b5e106708021..26dbfcb2a4aa 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Config.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Config.c @@ -32,6 +32,7 @@ static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______ma static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__47; static lean_object* l_Lean_Elab_Tactic_configElab___closed__17; static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__224; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__232; static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__210; static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__150; static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__69; @@ -634,7 +635,7 @@ static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config__ _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("def", 3); +x_1 = lean_mk_string_from_bytes("definition", 10); return x_1; } } @@ -654,23 +655,31 @@ static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config__ _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("declId", 6); +x_1 = lean_mk_string_from_bytes("def", 3); return x_1; } } static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__15() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("declId", 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___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_Elab_Tactic_configElab___closed__1; x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__3; x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__4; -x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__14; +x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__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_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__16() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__17() { _start: { lean_object* x_1; @@ -678,26 +687,26 @@ x_1 = lean_mk_string_from_bytes("evalUnsafe", 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__17() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__18() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__16; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__17; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__18() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__16; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__17; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__19() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__20() { _start: { lean_object* x_1; @@ -705,19 +714,19 @@ x_1 = lean_mk_string_from_bytes("optDeclSig", 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__20() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___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_Elab_Tactic_configElab___closed__1; x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__3; x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__4; -x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__19; +x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___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_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__21() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__22() { _start: { lean_object* x_1; @@ -725,7 +734,7 @@ x_1 = lean_mk_string_from_bytes("Term", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__22() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__23() { _start: { lean_object* x_1; @@ -733,19 +742,19 @@ x_1 = lean_mk_string_from_bytes("explicitBinder", 14); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__23() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__24() { _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_Tactic_configElab___closed__1; x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__3; -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__21; -x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__22; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__22; +x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__23; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__24() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__25() { _start: { lean_object* x_1; @@ -753,7 +762,7 @@ x_1 = lean_mk_string_from_bytes("(", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__25() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__26() { _start: { lean_object* x_1; @@ -761,26 +770,26 @@ x_1 = lean_mk_string_from_bytes("e", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__26() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__27() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__25; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__26; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__27() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__28() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__25; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__26; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__28() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__29() { _start: { lean_object* x_1; @@ -788,7 +797,7 @@ x_1 = lean_mk_string_from_bytes(":", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__29() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__30() { _start: { lean_object* x_1; @@ -796,82 +805,82 @@ x_1 = lean_mk_string_from_bytes("Expr", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__30() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__31() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__29; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__30; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__31() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__32() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__29; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__30; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__32() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__33() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Tactic_configElab___closed__1; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__29; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__30; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__33() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__34() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__32; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__33; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__34() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__35() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__32; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__33; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__35() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__36() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__34; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__35; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__36() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__37() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__33; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__35; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__34; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__36; 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_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__37() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__38() { _start: { lean_object* x_1; @@ -879,7 +888,7 @@ x_1 = lean_mk_string_from_bytes(")", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__38() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__39() { _start: { lean_object* x_1; @@ -887,19 +896,19 @@ x_1 = lean_mk_string_from_bytes("typeSpec", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__39() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__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_Lean_Elab_Tactic_configElab___closed__1; x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__3; -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__21; -x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__38; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__22; +x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__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_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__40() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__41() { _start: { lean_object* x_1; @@ -907,19 +916,19 @@ x_1 = lean_mk_string_from_bytes("app", 3); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__41() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__42() { _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_Tactic_configElab___closed__1; x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__3; -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__21; -x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__40; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__22; +x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__41; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__42() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__43() { _start: { lean_object* x_1; @@ -927,62 +936,62 @@ x_1 = lean_mk_string_from_bytes("TermElabM", 9); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__43() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__44() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__42; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__43; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__44() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__45() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__42; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__43; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__45() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__46() { _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_Tactic_configElab___closed__1; x_2 = l_Lean_Elab_Tactic_configElab___closed__2; -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__21; -x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__42; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__22; +x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___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_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__46() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__47() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__45; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__46; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__47() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__48() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__46; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__47; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__48() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__49() { _start: { lean_object* x_1; @@ -990,19 +999,19 @@ x_1 = lean_mk_string_from_bytes("declValSimple", 13); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__49() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___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_Tactic_configElab___closed__1; x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__3; x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__4; -x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__48; +x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___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_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__50() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__51() { _start: { lean_object* x_1; @@ -1010,7 +1019,7 @@ x_1 = lean_mk_string_from_bytes(":=", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__51() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__52() { _start: { lean_object* x_1; @@ -1018,16 +1027,16 @@ x_1 = lean_mk_string_from_bytes("Meta.evalExpr'", 14); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__52() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__53() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__51; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__52; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__53() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__54() { _start: { lean_object* x_1; @@ -1035,7 +1044,7 @@ x_1 = lean_mk_string_from_bytes("Meta", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__54() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__55() { _start: { lean_object* x_1; @@ -1043,74 +1052,74 @@ x_1 = lean_mk_string_from_bytes("evalExpr'", 9); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__55() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__56() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__53; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__54; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__54; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__55; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__56() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__57() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Elab_Tactic_configElab___closed__1; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__53; -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__54; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__54; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__55; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__57() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__58() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__56; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__57; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__58() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__59() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__56; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__57; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__59() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__60() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__58; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__59; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__60() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__61() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__57; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__59; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__58; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__60; 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_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__61() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__62() { _start: { lean_object* x_1; @@ -1118,19 +1127,19 @@ x_1 = lean_mk_string_from_bytes("namedArgument", 13); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__62() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___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_Tactic_configElab___closed__1; x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__3; -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__21; -x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__61; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__22; +x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__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_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__63() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__64() { _start: { lean_object* x_1; @@ -1138,26 +1147,26 @@ x_1 = lean_mk_string_from_bytes("safety", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__64() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__65() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__63; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__64; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__65() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__66() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__63; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__64; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__66() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__67() { _start: { lean_object* x_1; @@ -1165,19 +1174,19 @@ x_1 = lean_mk_string_from_bytes("dotIdent", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__67() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__68() { _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_Tactic_configElab___closed__1; x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__3; -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__21; -x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__66; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__22; +x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__67; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__68() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__69() { _start: { lean_object* x_1; @@ -1185,7 +1194,7 @@ x_1 = lean_mk_string_from_bytes(".", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__69() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__70() { _start: { lean_object* x_1; lean_object* x_2; @@ -1194,7 +1203,7 @@ x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__70() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__71() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -1204,7 +1213,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__71() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__72() { _start: { lean_object* x_1; @@ -1212,19 +1221,19 @@ x_1 = lean_mk_string_from_bytes("doubleQuotedName", 16); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__72() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__73() { _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_Tactic_configElab___closed__1; x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__3; -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__21; -x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__71; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__22; +x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__72; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__73() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__74() { _start: { lean_object* x_1; @@ -1232,7 +1241,7 @@ x_1 = lean_mk_string_from_bytes("`", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__74() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__75() { _start: { lean_object* x_1; @@ -1240,7 +1249,7 @@ x_1 = lean_mk_string_from_bytes("Termination", 11); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__75() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__76() { _start: { lean_object* x_1; @@ -1248,19 +1257,19 @@ x_1 = lean_mk_string_from_bytes("suffix", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__76() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__77() { _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_Tactic_configElab___closed__1; x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__3; -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__74; -x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__75; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__75; +x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__76; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__77() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__78() { _start: { lean_object* x_1; @@ -1268,19 +1277,19 @@ x_1 = lean_mk_string_from_bytes("attributes", 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__78() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__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_Lean_Elab_Tactic_configElab___closed__1; x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__3; -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__21; -x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__77; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__22; +x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__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_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__79() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__80() { _start: { lean_object* x_1; @@ -1288,7 +1297,7 @@ x_1 = lean_mk_string_from_bytes("@[", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__80() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__81() { _start: { lean_object* x_1; @@ -1296,19 +1305,19 @@ x_1 = lean_mk_string_from_bytes("attrInstance", 12); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__81() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__82() { _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_Tactic_configElab___closed__1; x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__3; -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__21; -x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__80; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__22; +x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__81; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__82() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__83() { _start: { lean_object* x_1; @@ -1316,19 +1325,19 @@ x_1 = lean_mk_string_from_bytes("attrKind", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__83() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__84() { _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_Tactic_configElab___closed__1; x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__3; -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__21; -x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__82; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__22; +x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__83; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__84() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__85() { _start: { lean_object* x_1; @@ -1336,7 +1345,7 @@ x_1 = lean_mk_string_from_bytes("Attr", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__85() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__86() { _start: { lean_object* x_1; @@ -1344,19 +1353,19 @@ x_1 = lean_mk_string_from_bytes("simple", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__86() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__87() { _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_Tactic_configElab___closed__1; x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__3; -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__84; -x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__85; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__85; +x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__86; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__87() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__88() { _start: { lean_object* x_1; @@ -1364,26 +1373,26 @@ x_1 = lean_mk_string_from_bytes("implemented_by", 14); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__88() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__89() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__87; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__88; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__89() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__90() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__87; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__88; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__90() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__91() { _start: { lean_object* x_1; @@ -1391,7 +1400,7 @@ x_1 = lean_mk_string_from_bytes("]", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__91() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__92() { _start: { lean_object* x_1; @@ -1399,19 +1408,19 @@ x_1 = lean_mk_string_from_bytes("opaque", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__92() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__93() { _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_Tactic_configElab___closed__1; x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__3; x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__4; -x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__91; +x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__92; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__93() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__94() { _start: { lean_object* x_1; @@ -1419,26 +1428,26 @@ x_1 = lean_mk_string_from_bytes("eval", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__94() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__95() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__93; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__94; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__95() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__96() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__93; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__94; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__96() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__97() { _start: { lean_object* x_1; @@ -1446,19 +1455,19 @@ x_1 = lean_mk_string_from_bytes("declSig", 7); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__97() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__98() { _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_Tactic_configElab___closed__1; x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__3; x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__4; -x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__96; +x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__97; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__98() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__99() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; @@ -1472,7 +1481,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__99() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__100() { _start: { lean_object* x_1; lean_object* x_2; @@ -1481,7 +1490,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__100() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__101() { _start: { lean_object* x_1; @@ -1489,26 +1498,26 @@ x_1 = lean_mk_string_from_bytes("optConfig", 9); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__101() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__102() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__100; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__101; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__102() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__103() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__100; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__101; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__103() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__104() { _start: { lean_object* x_1; @@ -1516,82 +1525,82 @@ x_1 = lean_mk_string_from_bytes("Syntax", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__104() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__105() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__103; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__104; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__105() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__106() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__103; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__104; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__106() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__107() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Tactic_configElab___closed__1; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__103; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__104; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__107() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__108() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__106; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__107; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__108() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__109() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__106; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__107; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__109() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__110() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__108; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__109; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__110() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__111() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__107; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__109; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__108; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__110; 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_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__111() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__112() { _start: { lean_object* x_1; @@ -1599,19 +1608,19 @@ x_1 = lean_mk_string_from_bytes("do", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__112() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__113() { _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_Tactic_configElab___closed__1; x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__3; -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__21; -x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__111; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__22; +x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__112; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__113() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__114() { _start: { lean_object* x_1; @@ -1619,19 +1628,19 @@ x_1 = lean_mk_string_from_bytes("doSeqIndent", 11); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__114() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__115() { _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_Tactic_configElab___closed__1; x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__3; -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__21; -x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__113; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__22; +x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__114; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__115() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__116() { _start: { lean_object* x_1; @@ -1639,19 +1648,19 @@ x_1 = lean_mk_string_from_bytes("doSeqItem", 9); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__116() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__117() { _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_Tactic_configElab___closed__1; x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__3; -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__21; -x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__115; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__22; +x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__116; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__117() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__118() { _start: { lean_object* x_1; @@ -1659,19 +1668,19 @@ x_1 = lean_mk_string_from_bytes("doIf", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__118() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__119() { _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_Tactic_configElab___closed__1; x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__3; -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__21; -x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__117; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__22; +x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__118; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__119() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__120() { _start: { lean_object* x_1; @@ -1679,7 +1688,7 @@ x_1 = lean_mk_string_from_bytes("if", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__120() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__121() { _start: { lean_object* x_1; @@ -1687,19 +1696,19 @@ x_1 = lean_mk_string_from_bytes("doIfProp", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__121() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__122() { _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_Tactic_configElab___closed__1; x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__3; -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__21; -x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__120; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__22; +x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__121; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__122() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__123() { _start: { lean_object* x_1; @@ -1707,16 +1716,16 @@ x_1 = lean_mk_string_from_bytes("optConfig.isNone", 16); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__123() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__124() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__122; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__123; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__124() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__125() { _start: { lean_object* x_1; @@ -1724,17 +1733,17 @@ x_1 = lean_mk_string_from_bytes("isNone", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__125() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__126() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__100; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__124; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__101; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__125; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__126() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__127() { _start: { lean_object* x_1; @@ -1742,7 +1751,7 @@ x_1 = lean_mk_string_from_bytes("then", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__127() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__128() { _start: { lean_object* x_1; @@ -1750,19 +1759,19 @@ x_1 = lean_mk_string_from_bytes("doReturn", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__128() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__129() { _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_Tactic_configElab___closed__1; x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__3; -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__21; -x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__127; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__22; +x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__128; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__129() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__130() { _start: { lean_object* x_1; @@ -1770,7 +1779,7 @@ x_1 = lean_mk_string_from_bytes("return", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__130() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__131() { _start: { lean_object* x_1; @@ -1778,19 +1787,19 @@ x_1 = lean_mk_string_from_bytes("structInst", 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__131() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__132() { _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_Tactic_configElab___closed__1; x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__3; -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__21; -x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__130; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__22; +x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__131; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__132() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__133() { _start: { lean_object* x_1; @@ -1798,7 +1807,7 @@ x_1 = lean_mk_string_from_bytes("{", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__133() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__134() { _start: { lean_object* x_1; @@ -1806,19 +1815,19 @@ x_1 = lean_mk_string_from_bytes("optEllipsis", 11); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__134() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__135() { _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_Tactic_configElab___closed__1; x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__3; -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__21; -x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__133; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__22; +x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__134; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__135() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__136() { _start: { lean_object* x_1; @@ -1826,7 +1835,7 @@ x_1 = lean_mk_string_from_bytes("}", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__136() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__137() { _start: { lean_object* x_1; @@ -1834,7 +1843,7 @@ x_1 = lean_mk_string_from_bytes("else", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__137() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__138() { _start: { lean_object* x_1; @@ -1842,19 +1851,19 @@ x_1 = lean_mk_string_from_bytes("doLetArrow", 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__138() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__139() { _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_Tactic_configElab___closed__1; x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__3; -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__21; -x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__137; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__22; +x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__138; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__139() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__140() { _start: { lean_object* x_1; @@ -1862,7 +1871,7 @@ x_1 = lean_mk_string_from_bytes("let", 3); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__140() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__141() { _start: { lean_object* x_1; @@ -1870,19 +1879,19 @@ x_1 = lean_mk_string_from_bytes("doIdDecl", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__141() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__142() { _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_Tactic_configElab___closed__1; x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__3; -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__21; -x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__140; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__22; +x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__141; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__142() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__143() { _start: { lean_object* x_1; @@ -1890,26 +1899,26 @@ x_1 = lean_mk_string_from_bytes("c", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__143() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__144() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__142; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__143; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__144() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__145() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__142; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__143; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__145() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__146() { _start: { lean_object* x_1; @@ -1917,7 +1926,7 @@ x_1 = lean_mk_string_from_bytes("←", 3); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__146() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__147() { _start: { lean_object* x_1; @@ -1925,19 +1934,19 @@ x_1 = lean_mk_string_from_bytes("doExpr", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__147() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__148() { _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_Tactic_configElab___closed__1; x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__3; -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__21; -x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__146; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__22; +x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__147; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__148() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__149() { _start: { lean_object* x_1; @@ -1945,17 +1954,17 @@ x_1 = lean_mk_string_from_bytes("term_<|_", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__149() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__150() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__148; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__149; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__150() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__151() { _start: { lean_object* x_1; @@ -1963,83 +1972,83 @@ x_1 = lean_mk_string_from_bytes("withoutModifyingStateWithInfoAndMessages", 40); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__151() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__152() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__150; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__151; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__152() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__153() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__150; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__151; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__153() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__154() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Elab_Tactic_configElab___closed__1; x_2 = l_Lean_Elab_Tactic_configElab___closed__2; -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__150; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__151; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__154() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__155() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__153; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__154; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__155() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__156() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__153; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__154; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__156() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__157() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__155; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__156; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__157() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__158() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__154; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__156; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__155; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__157; 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_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__158() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__159() { _start: { lean_object* x_1; @@ -2047,7 +2056,7 @@ x_1 = lean_mk_string_from_bytes("<|", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__159() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__160() { _start: { lean_object* x_1; @@ -2055,83 +2064,83 @@ x_1 = lean_mk_string_from_bytes("withLCtx", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__160() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__161() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__159; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__160; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__161() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__162() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__159; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__160; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__162() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__163() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Elab_Tactic_configElab___closed__1; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__53; -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__159; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__54; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__160; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__163() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__164() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__162; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__163; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__164() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__165() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__162; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__163; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__165() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__166() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__164; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__165; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__166() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__167() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__163; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__165; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__164; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__166; 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_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__167() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__168() { _start: { lean_object* x_1; @@ -2139,17 +2148,17 @@ x_1 = lean_mk_string_from_bytes("choice", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__168() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__169() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__167; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__168; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__169() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__170() { _start: { lean_object* x_1; @@ -2157,17 +2166,17 @@ x_1 = lean_mk_string_from_bytes("term{}", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__170() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__171() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__169; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__170; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__171() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__172() { _start: { lean_object* x_1; @@ -2175,83 +2184,83 @@ x_1 = lean_mk_string_from_bytes("withSaveInfoContext", 19); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__172() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__173() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__171; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__172; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__173() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__174() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__171; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__172; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__174() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__175() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Elab_Tactic_configElab___closed__1; x_2 = l_Lean_Elab_Tactic_configElab___closed__2; -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__171; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__172; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__175() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__176() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__174; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__175; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__176() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__177() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__174; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__175; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__177() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__178() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__176; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__177; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__178() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__179() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__175; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__177; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__176; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__178; 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_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__179() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__180() { _start: { lean_object* x_1; @@ -2259,16 +2268,16 @@ x_1 = lean_mk_string_from_bytes("Term.withSynthesize", 19); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__180() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__181() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__179; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__180; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__181() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__182() { _start: { lean_object* x_1; @@ -2276,75 +2285,75 @@ x_1 = lean_mk_string_from_bytes("withSynthesize", 14); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__182() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__183() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__21; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__181; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__22; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__182; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__183() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__184() { _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_Tactic_configElab___closed__1; x_2 = l_Lean_Elab_Tactic_configElab___closed__2; -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__21; -x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__181; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__22; +x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__182; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__184() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__185() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__183; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__184; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__185() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__186() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__183; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__184; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__186() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__187() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__185; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__186; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__187() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__188() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__184; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__186; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__185; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__187; 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_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__188() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__189() { _start: { lean_object* x_1; @@ -2352,16 +2361,16 @@ x_1 = lean_mk_string_from_bytes("Term.elabTermEnsuringType", 25); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__189() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__190() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__188; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__189; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__190() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__191() { _start: { lean_object* x_1; @@ -2369,75 +2378,75 @@ x_1 = lean_mk_string_from_bytes("elabTermEnsuringType", 20); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__191() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__192() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__21; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__190; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__22; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__191; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__192() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__193() { _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_Tactic_configElab___closed__1; x_2 = l_Lean_Elab_Tactic_configElab___closed__2; -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__21; -x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__190; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__22; +x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__191; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__193() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__194() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__192; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__193; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__194() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__195() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__192; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__193; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__195() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__196() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__194; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__195; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__196() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__197() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__193; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__195; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__194; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__196; 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_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__197() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__198() { _start: { lean_object* x_1; @@ -2445,17 +2454,17 @@ x_1 = lean_mk_string_from_bytes("term__[_]", 9); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__198() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__199() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__197; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__198; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__199() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__200() { _start: { lean_object* x_1; @@ -2463,7 +2472,7 @@ x_1 = lean_mk_string_from_bytes("[", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__200() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__201() { _start: { lean_object* x_1; @@ -2471,17 +2480,17 @@ x_1 = lean_mk_string_from_bytes("num", 3); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__201() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__202() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__200; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__201; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__202() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__203() { _start: { lean_object* x_1; @@ -2489,7 +2498,7 @@ x_1 = lean_mk_string_from_bytes("0", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__203() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__204() { _start: { lean_object* x_1; @@ -2497,7 +2506,7 @@ x_1 = lean_mk_string_from_bytes("3", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__204() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__205() { _start: { lean_object* x_1; @@ -2505,19 +2514,19 @@ x_1 = lean_mk_string_from_bytes("paren", 5); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__205() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__206() { _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_Tactic_configElab___closed__1; x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__3; -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__21; -x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__204; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__22; +x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__205; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__206() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__207() { _start: { lean_object* x_1; @@ -2525,16 +2534,16 @@ x_1 = lean_mk_string_from_bytes("Lean.mkConst", 12); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__207() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__208() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__206; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__207; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__208() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__209() { _start: { lean_object* x_1; @@ -2542,63 +2551,63 @@ x_1 = lean_mk_string_from_bytes("mkConst", 7); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__209() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__210() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Tactic_configElab___closed__1; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__208; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__209; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__210() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__211() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__209; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__210; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__211() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__212() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__209; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__210; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__212() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__213() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__211; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__212; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__213() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__214() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__210; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__212; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__211; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__213; 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_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__214() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__215() { _start: { lean_object* x_1; @@ -2606,16 +2615,16 @@ x_1 = lean_mk_string_from_bytes("Term.synthesizeSyntheticMVarsNoPostponing", 41) return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__215() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__216() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__214; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__215; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__216() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__217() { _start: { lean_object* x_1; @@ -2623,75 +2632,75 @@ x_1 = lean_mk_string_from_bytes("synthesizeSyntheticMVarsNoPostponing", 36); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__217() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__218() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__21; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__216; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__22; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__217; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__218() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__219() { _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_Tactic_configElab___closed__1; x_2 = l_Lean_Elab_Tactic_configElab___closed__2; -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__21; -x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__216; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__22; +x_4 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__217; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__219() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__220() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__218; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__219; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__220() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__221() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__218; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__219; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__221() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__222() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__220; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__221; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__222() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__223() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__219; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__221; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__220; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__222; 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_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__223() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__224() { _start: { lean_object* x_1; @@ -2699,82 +2708,82 @@ x_1 = lean_mk_string_from_bytes("instantiateMVars", 16); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__224() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__225() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__223; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__224; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__225() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__226() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__223; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__224; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__226() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__227() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Tactic_configElab___closed__1; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__223; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__224; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__227() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__228() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__226; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__227; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__228() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__229() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__226; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__227; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__229() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__230() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__228; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__229; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__230() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__231() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__227; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__229; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__228; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__230; 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_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__231() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__232() { _start: { lean_object* x_1; lean_object* x_2; @@ -2874,38 +2883,38 @@ x_28 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lea lean_inc_n(x_22, 5); lean_inc(x_17); x_29 = l_Lean_Syntax_node6(x_17, x_28, x_22, x_22, x_22, x_22, x_27, x_22); -x_30 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__12; +x_30 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__14; lean_inc(x_17); x_31 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_31, 0, x_17); lean_ctor_set(x_31, 1, x_30); -x_32 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__18; +x_32 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__19; lean_inc(x_18); lean_inc(x_19); x_33 = l_Lean_addMacroScope(x_19, x_32, x_18); x_34 = lean_box(0); -x_35 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__17; +x_35 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__18; lean_inc(x_17); x_36 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_36, 0, x_17); lean_ctor_set(x_36, 1, x_35); lean_ctor_set(x_36, 2, x_33); lean_ctor_set(x_36, 3, x_34); -x_37 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__15; +x_37 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__16; lean_inc(x_22); lean_inc(x_36); lean_inc(x_17); x_38 = l_Lean_Syntax_node2(x_17, x_37, x_36, x_22); -x_39 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__24; +x_39 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__25; lean_inc(x_17); x_40 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_40, 0, x_17); lean_ctor_set(x_40, 1, x_39); -x_41 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__27; +x_41 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__28; lean_inc(x_18); lean_inc(x_19); x_42 = l_Lean_addMacroScope(x_19, x_41, x_18); -x_43 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__26; +x_43 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__27; lean_inc(x_17); x_44 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_44, 0, x_17); @@ -2915,17 +2924,17 @@ lean_ctor_set(x_44, 3, x_34); lean_inc(x_44); lean_inc(x_17); x_45 = l_Lean_Syntax_node1(x_17, x_20, x_44); -x_46 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__28; +x_46 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__29; lean_inc(x_17); x_47 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_47, 0, x_17); lean_ctor_set(x_47, 1, x_46); -x_48 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__31; +x_48 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__32; lean_inc(x_18); lean_inc(x_19); x_49 = l_Lean_addMacroScope(x_19, x_48, x_18); -x_50 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__30; -x_51 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__36; +x_50 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__31; +x_51 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__37; lean_inc(x_17); x_52 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_52, 0, x_17); @@ -2935,12 +2944,12 @@ lean_ctor_set(x_52, 3, x_51); lean_inc(x_47); lean_inc(x_17); x_53 = l_Lean_Syntax_node2(x_17, x_20, x_47, x_52); -x_54 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__37; +x_54 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__38; lean_inc(x_17); x_55 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_55, 0, x_17); lean_ctor_set(x_55, 1, x_54); -x_56 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__23; +x_56 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__24; lean_inc(x_55); lean_inc(x_22); lean_inc(x_40); @@ -2948,12 +2957,12 @@ lean_inc(x_17); x_57 = l_Lean_Syntax_node5(x_17, x_56, x_40, x_45, x_53, x_22, x_55); lean_inc(x_17); x_58 = l_Lean_Syntax_node1(x_17, x_20, x_57); -x_59 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__44; +x_59 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__45; lean_inc(x_18); lean_inc(x_19); x_60 = l_Lean_addMacroScope(x_19, x_59, x_18); -x_61 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__43; -x_62 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__47; +x_61 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__44; +x_62 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__48; lean_inc(x_17); x_63 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_63, 0, x_17); @@ -2963,80 +2972,80 @@ lean_ctor_set(x_63, 3, x_62); lean_inc(x_13); lean_inc(x_17); x_64 = l_Lean_Syntax_node1(x_17, x_20, x_13); -x_65 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__41; +x_65 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__42; lean_inc(x_17); x_66 = l_Lean_Syntax_node2(x_17, x_65, x_63, x_64); -x_67 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__39; +x_67 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__40; lean_inc(x_47); lean_inc(x_17); x_68 = l_Lean_Syntax_node2(x_17, x_67, x_47, x_66); lean_inc(x_68); lean_inc(x_17); x_69 = l_Lean_Syntax_node1(x_17, x_20, x_68); -x_70 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__20; +x_70 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__21; lean_inc(x_69); lean_inc(x_58); lean_inc(x_17); x_71 = l_Lean_Syntax_node2(x_17, x_70, x_58, x_69); -x_72 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__50; +x_72 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__51; lean_inc(x_17); x_73 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_73, 0, x_17); lean_ctor_set(x_73, 1, x_72); -x_74 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__55; +x_74 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__56; lean_inc(x_18); lean_inc(x_19); x_75 = l_Lean_addMacroScope(x_19, x_74, x_18); -x_76 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__52; -x_77 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__60; +x_76 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__53; +x_77 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__61; lean_inc(x_17); x_78 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_78, 0, x_17); lean_ctor_set(x_78, 1, x_76); lean_ctor_set(x_78, 2, x_75); lean_ctor_set(x_78, 3, x_77); -x_79 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__65; +x_79 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__66; lean_inc(x_18); lean_inc(x_19); x_80 = l_Lean_addMacroScope(x_19, x_79, x_18); -x_81 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__64; +x_81 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__65; lean_inc(x_17); x_82 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_82, 0, x_17); lean_ctor_set(x_82, 1, x_81); lean_ctor_set(x_82, 2, x_80); lean_ctor_set(x_82, 3, x_34); -x_83 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__68; +x_83 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__69; lean_inc(x_17); x_84 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_84, 0, x_17); lean_ctor_set(x_84, 1, x_83); -x_85 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__70; +x_85 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__71; lean_inc(x_18); lean_inc(x_19); x_86 = l_Lean_addMacroScope(x_19, x_85, x_18); -x_87 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__69; +x_87 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__70; lean_inc(x_17); x_88 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_88, 0, x_17); lean_ctor_set(x_88, 1, x_87); lean_ctor_set(x_88, 2, x_86); lean_ctor_set(x_88, 3, x_34); -x_89 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__67; +x_89 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__68; lean_inc(x_17); x_90 = l_Lean_Syntax_node2(x_17, x_89, x_84, x_88); -x_91 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__62; +x_91 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__63; lean_inc(x_55); lean_inc(x_73); lean_inc(x_40); lean_inc(x_17); x_92 = l_Lean_Syntax_node5(x_17, x_91, x_40, x_82, x_73, x_90, x_55); -x_93 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__73; +x_93 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__74; lean_inc(x_17); x_94 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_94, 0, x_17); lean_ctor_set(x_94, 1, x_93); -x_95 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__72; +x_95 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__73; lean_inc(x_13); lean_inc(x_94); lean_inc(x_17); @@ -3047,11 +3056,11 @@ lean_inc(x_17); x_97 = l_Lean_Syntax_node4(x_17, x_20, x_92, x_13, x_96, x_44); lean_inc(x_17); x_98 = l_Lean_Syntax_node2(x_17, x_65, x_78, x_97); -x_99 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__76; +x_99 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__77; lean_inc_n(x_22, 2); lean_inc(x_17); x_100 = l_Lean_Syntax_node2(x_17, x_99, x_22, x_22); -x_101 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__49; +x_101 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__50; lean_inc(x_22); lean_inc(x_100); lean_inc(x_73); @@ -3065,20 +3074,20 @@ x_104 = l_Lean_Syntax_node5(x_17, x_103, x_31, x_38, x_71, x_102, x_22); x_105 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__6; lean_inc(x_17); x_106 = l_Lean_Syntax_node2(x_17, x_105, x_29, x_104); -x_107 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__79; +x_107 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__80; lean_inc(x_17); x_108 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_108, 0, x_17); lean_ctor_set(x_108, 1, x_107); -x_109 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__83; +x_109 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__84; lean_inc(x_22); lean_inc(x_17); x_110 = l_Lean_Syntax_node1(x_17, x_109, x_22); -x_111 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__89; +x_111 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__90; lean_inc(x_18); lean_inc(x_19); x_112 = l_Lean_addMacroScope(x_19, x_111, x_18); -x_113 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__88; +x_113 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__89; lean_inc(x_17); x_114 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_114, 0, x_17); @@ -3087,20 +3096,20 @@ lean_ctor_set(x_114, 2, x_112); lean_ctor_set(x_114, 3, x_34); lean_inc(x_17); x_115 = l_Lean_Syntax_node1(x_17, x_20, x_36); -x_116 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__86; +x_116 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__87; lean_inc(x_17); x_117 = l_Lean_Syntax_node2(x_17, x_116, x_114, x_115); -x_118 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__81; +x_118 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__82; lean_inc(x_17); x_119 = l_Lean_Syntax_node2(x_17, x_118, x_110, x_117); lean_inc(x_17); x_120 = l_Lean_Syntax_node1(x_17, x_20, x_119); -x_121 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__90; +x_121 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__91; lean_inc(x_17); x_122 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_122, 0, x_17); lean_ctor_set(x_122, 1, x_121); -x_123 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__78; +x_123 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__79; lean_inc(x_122); lean_inc(x_17); x_124 = l_Lean_Syntax_node3(x_17, x_123, x_108, x_120, x_122); @@ -3109,16 +3118,16 @@ x_125 = l_Lean_Syntax_node1(x_17, x_20, x_124); lean_inc_n(x_22, 5); lean_inc(x_17); x_126 = l_Lean_Syntax_node6(x_17, x_28, x_22, x_125, x_22, x_22, x_22, x_22); -x_127 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__91; +x_127 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__92; lean_inc(x_17); x_128 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_128, 0, x_17); lean_ctor_set(x_128, 1, x_127); -x_129 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__95; +x_129 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__96; lean_inc(x_18); lean_inc(x_19); x_130 = l_Lean_addMacroScope(x_19, x_129, x_18); -x_131 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__94; +x_131 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__95; lean_inc(x_17); x_132 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_132, 0, x_17); @@ -3129,29 +3138,29 @@ lean_inc(x_22); lean_inc(x_132); lean_inc(x_17); x_133 = l_Lean_Syntax_node2(x_17, x_37, x_132, x_22); -x_134 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__97; +x_134 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__98; lean_inc(x_17); x_135 = l_Lean_Syntax_node2(x_17, x_134, x_58, x_68); -x_136 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__92; +x_136 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__93; lean_inc(x_22); lean_inc(x_17); x_137 = l_Lean_Syntax_node4(x_17, x_136, x_128, x_133, x_135, x_22); lean_inc(x_17); x_138 = l_Lean_Syntax_node2(x_17, x_105, x_126, x_137); -x_139 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__99; +x_139 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__100; x_140 = lean_array_push(x_139, x_11); -x_141 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__98; +x_141 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__99; x_142 = lean_array_push(x_140, x_141); x_143 = lean_box(2); x_144 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_144, 0, x_143); lean_ctor_set(x_144, 1, x_37); lean_ctor_set(x_144, 2, x_142); -x_145 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__102; +x_145 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__103; lean_inc(x_18); lean_inc(x_19); x_146 = l_Lean_addMacroScope(x_19, x_145, x_18); -x_147 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__101; +x_147 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__102; lean_inc(x_17); x_148 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_148, 0, x_17); @@ -3161,12 +3170,12 @@ lean_ctor_set(x_148, 3, x_34); lean_inc(x_148); lean_inc(x_17); x_149 = l_Lean_Syntax_node1(x_17, x_20, x_148); -x_150 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__105; +x_150 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__106; lean_inc(x_18); lean_inc(x_19); x_151 = l_Lean_addMacroScope(x_19, x_150, x_18); -x_152 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__104; -x_153 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__110; +x_152 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__105; +x_153 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__111; lean_inc(x_17); x_154 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_154, 0, x_17); @@ -3185,58 +3194,58 @@ lean_inc(x_17); x_157 = l_Lean_Syntax_node1(x_17, x_20, x_156); lean_inc(x_17); x_158 = l_Lean_Syntax_node2(x_17, x_70, x_157, x_69); -x_159 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__111; +x_159 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__112; lean_inc(x_17); x_160 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_160, 0, x_17); lean_ctor_set(x_160, 1, x_159); -x_161 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__119; +x_161 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__120; lean_inc(x_17); x_162 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_162, 0, x_17); lean_ctor_set(x_162, 1, x_161); -x_163 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__125; +x_163 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__126; lean_inc(x_18); lean_inc(x_19); x_164 = l_Lean_addMacroScope(x_19, x_163, x_18); -x_165 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__123; +x_165 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__124; lean_inc(x_17); x_166 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_166, 0, x_17); lean_ctor_set(x_166, 1, x_165); lean_ctor_set(x_166, 2, x_164); lean_ctor_set(x_166, 3, x_34); -x_167 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__121; +x_167 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__122; lean_inc(x_22); lean_inc(x_17); x_168 = l_Lean_Syntax_node2(x_17, x_167, x_22, x_166); -x_169 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__126; +x_169 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__127; lean_inc(x_17); x_170 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_170, 0, x_17); lean_ctor_set(x_170, 1, x_169); -x_171 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__129; +x_171 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__130; lean_inc(x_17); x_172 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_172, 0, x_17); lean_ctor_set(x_172, 1, x_171); -x_173 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__132; +x_173 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__133; lean_inc(x_17); x_174 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_174, 0, x_17); lean_ctor_set(x_174, 1, x_173); -x_175 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__134; +x_175 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__135; lean_inc(x_22); lean_inc(x_17); x_176 = l_Lean_Syntax_node1(x_17, x_175, x_22); lean_inc(x_17); x_177 = l_Lean_Syntax_node2(x_17, x_20, x_47, x_13); -x_178 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__135; +x_178 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__136; lean_inc(x_17); x_179 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_179, 0, x_17); lean_ctor_set(x_179, 1, x_178); -x_180 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__131; +x_180 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__132; lean_inc(x_179); lean_inc(x_176); lean_inc_n(x_22, 2); @@ -3245,74 +3254,74 @@ lean_inc(x_17); x_181 = l_Lean_Syntax_node6(x_17, x_180, x_174, x_22, x_22, x_176, x_177, x_179); lean_inc(x_17); x_182 = l_Lean_Syntax_node1(x_17, x_20, x_181); -x_183 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__128; +x_183 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__129; lean_inc(x_17); x_184 = l_Lean_Syntax_node2(x_17, x_183, x_172, x_182); -x_185 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__116; +x_185 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__117; lean_inc(x_22); lean_inc(x_17); x_186 = l_Lean_Syntax_node2(x_17, x_185, x_184, x_22); lean_inc(x_17); x_187 = l_Lean_Syntax_node1(x_17, x_20, x_186); -x_188 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__114; +x_188 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__115; lean_inc(x_17); x_189 = l_Lean_Syntax_node1(x_17, x_188, x_187); -x_190 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__136; +x_190 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__137; lean_inc(x_17); x_191 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_191, 0, x_17); lean_ctor_set(x_191, 1, x_190); -x_192 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__139; +x_192 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__140; lean_inc(x_17); x_193 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_193, 0, x_17); lean_ctor_set(x_193, 1, x_192); -x_194 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__144; +x_194 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__145; lean_inc(x_18); lean_inc(x_19); x_195 = l_Lean_addMacroScope(x_19, x_194, x_18); -x_196 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__143; +x_196 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__144; lean_inc(x_17); x_197 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_197, 0, x_17); lean_ctor_set(x_197, 1, x_196); lean_ctor_set(x_197, 2, x_195); lean_ctor_set(x_197, 3, x_34); -x_198 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__145; +x_198 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__146; lean_inc(x_17); x_199 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_199, 0, x_17); lean_ctor_set(x_199, 1, x_198); -x_200 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__152; +x_200 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__153; lean_inc(x_18); lean_inc(x_19); x_201 = l_Lean_addMacroScope(x_19, x_200, x_18); -x_202 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__151; -x_203 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__157; +x_202 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__152; +x_203 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__158; lean_inc(x_17); x_204 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_204, 0, x_17); lean_ctor_set(x_204, 1, x_202); lean_ctor_set(x_204, 2, x_201); lean_ctor_set(x_204, 3, x_203); -x_205 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__158; +x_205 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__159; lean_inc(x_17); x_206 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_206, 0, x_17); lean_ctor_set(x_206, 1, x_205); -x_207 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__161; +x_207 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__162; lean_inc(x_18); lean_inc(x_19); x_208 = l_Lean_addMacroScope(x_19, x_207, x_18); -x_209 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__160; -x_210 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__166; +x_209 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__161; +x_210 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__167; lean_inc(x_17); x_211 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_211, 0, x_17); lean_ctor_set(x_211, 1, x_209); lean_ctor_set(x_211, 2, x_208); lean_ctor_set(x_211, 3, x_210); -x_212 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__170; +x_212 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__171; lean_inc(x_179); lean_inc(x_174); lean_inc(x_17); @@ -3320,7 +3329,7 @@ x_213 = l_Lean_Syntax_node2(x_17, x_212, x_174, x_179); lean_inc_n(x_22, 3); lean_inc(x_17); x_214 = l_Lean_Syntax_node6(x_17, x_180, x_174, x_22, x_22, x_176, x_22, x_179); -x_215 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__168; +x_215 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__169; lean_inc(x_17); x_216 = l_Lean_Syntax_node2(x_17, x_215, x_213, x_214); lean_inc(x_216); @@ -3328,61 +3337,61 @@ lean_inc(x_17); x_217 = l_Lean_Syntax_node2(x_17, x_20, x_216, x_216); lean_inc(x_17); x_218 = l_Lean_Syntax_node2(x_17, x_65, x_211, x_217); -x_219 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__173; +x_219 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__174; lean_inc(x_18); lean_inc(x_19); x_220 = l_Lean_addMacroScope(x_19, x_219, x_18); -x_221 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__172; -x_222 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__178; +x_221 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__173; +x_222 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__179; lean_inc(x_17); x_223 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_223, 0, x_17); lean_ctor_set(x_223, 1, x_221); lean_ctor_set(x_223, 2, x_220); lean_ctor_set(x_223, 3, x_222); -x_224 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__182; +x_224 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__183; lean_inc(x_18); lean_inc(x_19); x_225 = l_Lean_addMacroScope(x_19, x_224, x_18); -x_226 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__180; -x_227 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__187; +x_226 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__181; +x_227 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__188; lean_inc(x_17); x_228 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_228, 0, x_17); lean_ctor_set(x_228, 1, x_226); lean_ctor_set(x_228, 2, x_225); lean_ctor_set(x_228, 3, x_227); -x_229 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__191; +x_229 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__192; lean_inc(x_18); lean_inc(x_19); x_230 = l_Lean_addMacroScope(x_19, x_229, x_18); -x_231 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__189; -x_232 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__196; +x_231 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__190; +x_232 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__197; lean_inc(x_17); x_233 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_233, 0, x_17); lean_ctor_set(x_233, 1, x_231); lean_ctor_set(x_233, 2, x_230); lean_ctor_set(x_233, 3, x_232); -x_234 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__199; +x_234 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__200; lean_inc(x_17); x_235 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_235, 0, x_17); lean_ctor_set(x_235, 1, x_234); -x_236 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__202; +x_236 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__203; lean_inc(x_17); x_237 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_237, 0, x_17); lean_ctor_set(x_237, 1, x_236); -x_238 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__201; +x_238 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__202; lean_inc(x_17); x_239 = l_Lean_Syntax_node1(x_17, x_238, x_237); -x_240 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__198; +x_240 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__199; lean_inc(x_122); lean_inc(x_235); lean_inc(x_17); x_241 = l_Lean_Syntax_node4(x_17, x_240, x_148, x_235, x_239, x_122); -x_242 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__203; +x_242 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__204; lean_inc(x_17); x_243 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_243, 0, x_17); @@ -3391,12 +3400,12 @@ lean_inc(x_17); x_244 = l_Lean_Syntax_node1(x_17, x_238, x_243); lean_inc(x_17); x_245 = l_Lean_Syntax_node4(x_17, x_240, x_241, x_235, x_244, x_122); -x_246 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__209; +x_246 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__210; lean_inc(x_18); lean_inc(x_19); x_247 = l_Lean_addMacroScope(x_19, x_246, x_18); -x_248 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__207; -x_249 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__213; +x_248 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__208; +x_249 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__214; lean_inc(x_17); x_250 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_250, 0, x_17); @@ -3407,23 +3416,23 @@ lean_inc(x_17); x_251 = l_Lean_Syntax_node1(x_17, x_20, x_96); lean_inc(x_17); x_252 = l_Lean_Syntax_node2(x_17, x_65, x_250, x_251); -x_253 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__205; +x_253 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__206; lean_inc(x_17); x_254 = l_Lean_Syntax_node3(x_17, x_253, x_40, x_252, x_55); lean_inc(x_17); x_255 = l_Lean_Syntax_node2(x_17, x_20, x_245, x_254); lean_inc(x_17); x_256 = l_Lean_Syntax_node2(x_17, x_65, x_233, x_255); -x_257 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__147; +x_257 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__148; lean_inc(x_17); x_258 = l_Lean_Syntax_node1(x_17, x_257, x_256); -x_259 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__141; +x_259 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__142; lean_inc(x_199); lean_inc(x_22); lean_inc(x_197); lean_inc(x_17); x_260 = l_Lean_Syntax_node4(x_17, x_259, x_197, x_22, x_199, x_258); -x_261 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__138; +x_261 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__139; lean_inc(x_22); lean_inc(x_193); lean_inc(x_17); @@ -3431,12 +3440,12 @@ x_262 = l_Lean_Syntax_node3(x_17, x_261, x_193, x_22, x_260); lean_inc(x_22); lean_inc(x_17); x_263 = l_Lean_Syntax_node2(x_17, x_185, x_262, x_22); -x_264 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__217; +x_264 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__218; lean_inc(x_18); lean_inc(x_19); x_265 = l_Lean_addMacroScope(x_19, x_264, x_18); -x_266 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__215; -x_267 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__222; +x_266 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__216; +x_267 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__223; lean_inc(x_17); x_268 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_268, 0, x_17); @@ -3448,10 +3457,10 @@ x_269 = l_Lean_Syntax_node1(x_17, x_257, x_268); lean_inc(x_22); lean_inc(x_17); x_270 = l_Lean_Syntax_node2(x_17, x_185, x_269, x_22); -x_271 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__225; +x_271 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__226; x_272 = l_Lean_addMacroScope(x_19, x_271, x_18); -x_273 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__224; -x_274 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__230; +x_273 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__225; +x_274 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__231; lean_inc(x_17); x_275 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_275, 0, x_17); @@ -3473,7 +3482,7 @@ lean_inc(x_17); x_280 = l_Lean_Syntax_node3(x_17, x_20, x_263, x_270, x_279); lean_inc(x_17); x_281 = l_Lean_Syntax_node1(x_17, x_188, x_280); -x_282 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__112; +x_282 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__113; lean_inc(x_160); lean_inc(x_17); x_283 = l_Lean_Syntax_node2(x_17, x_282, x_160, x_281); @@ -3481,7 +3490,7 @@ lean_inc(x_17); x_284 = l_Lean_Syntax_node1(x_17, x_20, x_283); lean_inc(x_17); x_285 = l_Lean_Syntax_node2(x_17, x_65, x_228, x_284); -x_286 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__149; +x_286 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__150; lean_inc(x_206); lean_inc(x_17); x_287 = l_Lean_Syntax_node3(x_17, x_286, x_223, x_206, x_285); @@ -3514,7 +3523,7 @@ lean_inc(x_17); x_298 = l_Lean_Syntax_node1(x_17, x_188, x_297); lean_inc(x_17); x_299 = l_Lean_Syntax_node2(x_17, x_20, x_191, x_298); -x_300 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__118; +x_300 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__119; lean_inc(x_22); lean_inc(x_17); x_301 = l_Lean_Syntax_node6(x_17, x_300, x_162, x_168, x_170, x_189, x_22, x_299); @@ -3536,7 +3545,7 @@ x_307 = l_Lean_Syntax_node5(x_17, x_103, x_31, x_144, x_158, x_306, x_22); if (lean_obj_tag(x_14) == 0) { lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; -x_308 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__231; +x_308 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__232; lean_inc(x_17); x_309 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_309, 0, x_17); @@ -4164,6 +4173,8 @@ l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__230); l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__231 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__231(); lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__231); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__232 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__232(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__232); l_Lean_Elab_Tactic_checkConfigElab___closed__1 = _init_l_Lean_Elab_Tactic_checkConfigElab___closed__1(); lean_mark_persistent(l_Lean_Elab_Tactic_checkConfigElab___closed__1); l___regBuiltin_Lean_Elab_Tactic_checkConfigElab___closed__1 = _init_l___regBuiltin_Lean_Elab_Tactic_checkConfigElab___closed__1(); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Simp/BuiltinSimprocs/UInt.c b/stage0/stdlib/Lean/Meta/Tactic/Simp/BuiltinSimprocs/UInt.c index 9bf614fc5085..42d3c44793d4 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Simp/BuiltinSimprocs/UInt.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Simp/BuiltinSimprocs/UInt.c @@ -648,6 +648,7 @@ LEAN_EXPORT lean_object* l_UInt8_reduceLT(lean_object*, lean_object*, lean_objec static lean_object* l___regBuiltin_UInt32_reduceSub_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_11260____closed__1; static lean_object* l___regBuiltin_UInt16_reduceBNe_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_9210__10228____closed__11; static lean_object* l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__466; +static lean_object* l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__503; static lean_object* l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__448; static lean_object* l___regBuiltin_UInt8_reduceBNe_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_7772__8790____closed__13; static lean_object* l___regBuiltin_UInt64_reduceBEq_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_12086__13064____closed__2; @@ -2115,7 +2116,7 @@ static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UI _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("def", 3); +x_1 = lean_mk_string_from_bytes("definition", 10); return x_1; } } @@ -2135,23 +2136,31 @@ static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UI _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("declId", 6); +x_1 = lean_mk_string_from_bytes("def", 3); return x_1; } } static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__25() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("declId", 6); +return x_1; +} +} +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___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___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__13; x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__14; -x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__24; +x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__25; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__26() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__27() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; @@ -2165,7 +2174,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__27() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__28() { _start: { lean_object* x_1; lean_object* x_2; @@ -2174,33 +2183,33 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__28() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__29() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__27; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__28; x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__9; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__29() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__30() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__28; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__26; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__29; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__27; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__30() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__31() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__25; -x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__29; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__26; +x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__30; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -2208,7 +2217,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__31() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__32() { _start: { lean_object* x_1; @@ -2216,19 +2225,19 @@ x_1 = lean_mk_string_from_bytes("optDeclSig", 10); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__32() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___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___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__13; x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__14; -x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__31; +x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__32; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__33() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__34() { _start: { lean_object* x_1; @@ -2236,7 +2245,7 @@ x_1 = lean_mk_string_from_bytes("Term", 4); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__34() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__35() { _start: { lean_object* x_1; @@ -2244,19 +2253,19 @@ x_1 = lean_mk_string_from_bytes("explicitBinder", 14); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__35() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____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___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__13; -x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__33; -x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__34; +x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__34; +x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____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___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__36() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__37() { _start: { lean_object* x_1; @@ -2264,7 +2273,7 @@ x_1 = lean_mk_string_from_bytes("(", 1); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__37() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__38() { _start: { lean_object* x_1; @@ -2272,26 +2281,26 @@ x_1 = lean_mk_string_from_bytes("e", 1); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__38() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__39() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__37; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__38; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__39() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__40() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__37; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__38; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__40() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__41() { _start: { lean_object* x_1; @@ -2299,7 +2308,7 @@ x_1 = lean_mk_string_from_bytes(":", 1); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__41() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__42() { _start: { lean_object* x_1; @@ -2307,82 +2316,82 @@ x_1 = lean_mk_string_from_bytes("Expr", 4); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__42() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__43() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__41; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__42; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__43() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__44() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__41; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__42; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__44() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__45() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__41; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__42; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__45() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__46() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__44; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__45; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__46() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__47() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__44; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__45; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__47() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__48() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__46; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__47; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__48() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__49() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__45; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__47; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__46; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__48; 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___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__49() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__50() { _start: { lean_object* x_1; @@ -2390,7 +2399,7 @@ x_1 = lean_mk_string_from_bytes(")", 1); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__50() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__51() { _start: { lean_object* x_1; @@ -2398,19 +2407,19 @@ x_1 = lean_mk_string_from_bytes("typeSpec", 8); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__51() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__52() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__13; -x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__33; -x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__50; +x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__34; +x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__51; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__52() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__53() { _start: { lean_object* x_1; @@ -2418,19 +2427,19 @@ x_1 = lean_mk_string_from_bytes("app", 3); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__53() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__54() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__13; -x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__33; -x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__52; +x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__34; +x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__53; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__54() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__55() { _start: { lean_object* x_1; @@ -2438,26 +2447,26 @@ x_1 = lean_mk_string_from_bytes("SimpM", 5); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__55() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__56() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__54; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__55; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__56() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__57() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__54; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__55; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__57() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58() { _start: { lean_object* x_1; @@ -2465,7 +2474,7 @@ x_1 = lean_mk_string_from_bytes("Meta", 4); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__59() { _start: { lean_object* x_1; @@ -2473,43 +2482,43 @@ x_1 = lean_mk_string_from_bytes("Simp", 4); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__59() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__60() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__57; -x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; -x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__54; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; +x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__59; +x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__55; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__60() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__61() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__59; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__60; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__61() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__62() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__60; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__61; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__62() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__63() { _start: { lean_object* x_1; @@ -2517,19 +2526,19 @@ x_1 = lean_mk_string_from_bytes("paren", 5); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__63() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__64() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__13; -x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__33; -x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__62; +x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__34; +x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__63; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__64() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__65() { _start: { lean_object* x_1; @@ -2537,104 +2546,104 @@ x_1 = lean_mk_string_from_bytes("Option", 6); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__65() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__66() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__64; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__65; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__66() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__67() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__64; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__65; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__67() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__68() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__66; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__67; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__68() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__69() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__66; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__67; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__69() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__70() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__64; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__65; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__70() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__71() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__69; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__70; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__71() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__72() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__70; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__71; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__72() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__73() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__68; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__71; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__69; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__72; 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___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__73() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__74() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__67; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__72; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__68; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__73; 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___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__74() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__75() { _start: { lean_object* x_1; @@ -2642,19 +2651,19 @@ x_1 = lean_mk_string_from_bytes("declValSimple", 13); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__75() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__76() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__13; x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__14; -x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__74; +x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__75; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__76() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__77() { _start: { lean_object* x_1; @@ -2662,7 +2671,7 @@ x_1 = lean_mk_string_from_bytes(":=", 2); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__77() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__78() { _start: { lean_object* x_1; @@ -2670,19 +2679,19 @@ x_1 = lean_mk_string_from_bytes("do", 2); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__78() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____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___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__13; -x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__33; -x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__77; +x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__34; +x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____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___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__79() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__80() { _start: { lean_object* x_1; @@ -2690,19 +2699,19 @@ x_1 = lean_mk_string_from_bytes("doSeqIndent", 11); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__80() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__81() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__13; -x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__33; -x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__79; +x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__34; +x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__80; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__81() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__82() { _start: { lean_object* x_1; @@ -2710,19 +2719,19 @@ x_1 = lean_mk_string_from_bytes("doSeqItem", 9); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__82() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__83() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__13; -x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__33; -x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__81; +x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__34; +x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__82; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__83() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__84() { _start: { lean_object* x_1; @@ -2730,19 +2739,19 @@ x_1 = lean_mk_string_from_bytes("doLetArrow", 10); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__84() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__85() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__13; -x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__33; -x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__83; +x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__34; +x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__84; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__85() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__86() { _start: { lean_object* x_1; @@ -2750,7 +2759,7 @@ x_1 = lean_mk_string_from_bytes("let", 3); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__86() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__87() { _start: { lean_object* x_1; @@ -2758,19 +2767,19 @@ x_1 = lean_mk_string_from_bytes("doPatDecl", 9); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__87() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__88() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__13; -x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__33; -x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__86; +x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__34; +x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__87; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__88() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__89() { _start: { lean_object* x_1; @@ -2778,60 +2787,60 @@ x_1 = lean_mk_string_from_bytes("some", 4); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__89() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__90() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__88; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__89; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__90() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__91() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__88; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__89; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__91() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__92() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__64; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__88; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__65; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__89; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__92() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__93() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__91; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__92; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__93() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__94() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__92; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__93; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__94() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__95() { _start: { lean_object* x_1; @@ -2839,19 +2848,19 @@ x_1 = lean_mk_string_from_bytes("tuple", 5); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__95() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__96() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__13; -x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__33; -x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__94; +x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__34; +x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__95; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__96() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__97() { _start: { lean_object* x_1; @@ -2859,26 +2868,26 @@ x_1 = lean_mk_string_from_bytes("n", 1); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__97() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__98() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__96; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__97; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__98() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__99() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__96; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__97; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__99() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__100() { _start: { lean_object* x_1; @@ -2886,7 +2895,7 @@ x_1 = lean_mk_string_from_bytes(",", 1); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__100() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__101() { _start: { lean_object* x_1; @@ -2894,19 +2903,19 @@ x_1 = lean_mk_string_from_bytes("hole", 4); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__101() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__102() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__13; -x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__33; -x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__100; +x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__34; +x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__101; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__102() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__103() { _start: { lean_object* x_1; @@ -2914,7 +2923,7 @@ x_1 = lean_mk_string_from_bytes("_", 1); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__103() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__104() { _start: { lean_object* x_1; @@ -2922,7 +2931,7 @@ x_1 = lean_mk_string_from_bytes("←", 3); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__104() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__105() { _start: { lean_object* x_1; @@ -2930,19 +2939,19 @@ x_1 = lean_mk_string_from_bytes("doExpr", 6); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__105() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__106() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__13; -x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__33; -x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__104; +x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__34; +x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__105; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__106() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__107() { _start: { lean_object* x_1; @@ -2950,83 +2959,83 @@ x_1 = lean_mk_string_from_bytes("getOfNatValue\?", 14); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__107() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__108() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__106; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__107; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__108() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__109() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__106; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__107; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__109() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__110() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__57; -x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__106; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; +x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__107; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__110() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__111() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__109; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__110; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__111() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__112() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__109; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__110; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__112() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__113() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__111; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__112; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__113() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__114() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__110; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__112; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__111; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__113; 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___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__114() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__115() { _start: { lean_object* x_1; @@ -3034,7 +3043,7 @@ x_1 = lean_mk_string_from_bytes("|", 1); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__115() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__116() { _start: { lean_object* x_1; @@ -3042,19 +3051,19 @@ x_1 = lean_mk_string_from_bytes("doReturn", 8); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__116() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__117() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__13; -x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__33; -x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__115; +x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__34; +x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__116; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__117() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__118() { _start: { lean_object* x_1; @@ -3062,7 +3071,7 @@ x_1 = lean_mk_string_from_bytes("return", 6); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__118() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__119() { _start: { lean_object* x_1; @@ -3070,60 +3079,60 @@ x_1 = lean_mk_string_from_bytes("none", 4); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__119() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__120() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__118; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__119; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__120() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__121() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__118; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__119; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__121() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__122() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__64; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__118; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__65; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__119; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__122() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__123() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__121; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__122; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__123() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__124() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__122; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__123; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__124() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__125() { _start: { lean_object* x_1; @@ -3131,7 +3140,7 @@ x_1 = lean_mk_string_from_bytes("Termination", 11); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__125() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__126() { _start: { lean_object* x_1; @@ -3139,19 +3148,19 @@ x_1 = lean_mk_string_from_bytes("suffix", 6); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__126() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__127() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__13; -x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__124; -x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__125; +x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__125; +x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__126; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__127() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__128() { _start: { lean_object* x_1; @@ -3159,19 +3168,19 @@ x_1 = lean_mk_string_from_bytes("attributes", 10); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__128() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__129() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__13; -x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__33; -x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__127; +x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__34; +x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__128; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__129() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__130() { _start: { lean_object* x_1; @@ -3179,7 +3188,7 @@ x_1 = lean_mk_string_from_bytes("@[", 2); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__130() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__131() { _start: { lean_object* x_1; @@ -3187,19 +3196,19 @@ x_1 = lean_mk_string_from_bytes("attrInstance", 12); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__131() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__132() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__13; -x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__33; -x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__130; +x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__34; +x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__131; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__132() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__133() { _start: { lean_object* x_1; @@ -3207,19 +3216,19 @@ x_1 = lean_mk_string_from_bytes("attrKind", 8); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__133() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__134() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__13; -x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__33; -x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__132; +x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__34; +x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__133; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__134() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__135() { _start: { lean_object* x_1; @@ -3227,7 +3236,7 @@ x_1 = lean_mk_string_from_bytes("Attr", 4); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__135() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__136() { _start: { lean_object* x_1; @@ -3235,19 +3244,19 @@ x_1 = lean_mk_string_from_bytes("simple", 6); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__136() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__137() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__13; -x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__134; -x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__135; +x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__135; +x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__136; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__137() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__138() { _start: { lean_object* x_1; @@ -3255,72 +3264,72 @@ x_1 = lean_mk_string_from_bytes("inline", 6); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__138() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__139() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__137; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__138; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__139() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__140() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__137; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__138; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__140() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__141() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__139; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__140; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__141() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__142() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__139; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__140; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__142() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__143() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__141; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__142; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__143() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__144() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__140; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__142; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__141; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__143; 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___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__144() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__145() { _start: { lean_object* x_1; @@ -3328,7 +3337,7 @@ x_1 = lean_mk_string_from_bytes("]", 1); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__145() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__146() { _start: { lean_object* x_1; @@ -3336,26 +3345,26 @@ x_1 = lean_mk_string_from_bytes("reduceBin", 9); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__146() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__147() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__145; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__146; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__147() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__148() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__145; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__146; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__148() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__149() { _start: { lean_object* x_1; @@ -3363,26 +3372,26 @@ x_1 = lean_mk_string_from_bytes("declName", 8); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__149() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__150() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__148; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__149; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__150() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__151() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__148; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__149; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__151() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__152() { _start: { lean_object* x_1; @@ -3390,82 +3399,82 @@ x_1 = lean_mk_string_from_bytes("Name", 4); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__152() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__153() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__151; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__152; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__153() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__154() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__151; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__152; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__154() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__155() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__151; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__152; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__155() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__156() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__154; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__155; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__156() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__157() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__154; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___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___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__157() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__158() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__156; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__157; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__158() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__159() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__155; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__157; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__156; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__158; 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___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__159() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__160() { _start: { lean_object* x_1; @@ -3473,26 +3482,26 @@ x_1 = lean_mk_string_from_bytes("arity", 5); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__160() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__161() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__159; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__160; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__161() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__162() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__159; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__160; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__162() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__163() { _start: { lean_object* x_1; @@ -3500,72 +3509,72 @@ x_1 = lean_mk_string_from_bytes("Nat", 3); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__163() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__164() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__162; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__163; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__164() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__165() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__162; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__163; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__165() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__166() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__164; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__165; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__166() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__167() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__164; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__165; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__167() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__168() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__166; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__167; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__168() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__169() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__165; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__167; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__166; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__168; 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___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__169() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__170() { _start: { lean_object* x_1; @@ -3573,26 +3582,26 @@ x_1 = lean_mk_string_from_bytes("op", 2); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__170() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__171() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__169; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__170; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__171() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__172() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__169; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__170; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__172() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__173() { _start: { lean_object* x_1; @@ -3600,19 +3609,19 @@ x_1 = lean_mk_string_from_bytes("arrow", 5); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__173() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__174() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__13; -x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__33; -x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__172; +x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__34; +x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__173; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__174() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__175() { _start: { lean_object* x_1; @@ -3620,7 +3629,7 @@ x_1 = lean_mk_string_from_bytes("→", 3); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__175() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__176() { _start: { lean_object* x_1; @@ -3628,84 +3637,84 @@ x_1 = lean_mk_string_from_bytes("DStep", 5); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__176() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__177() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__175; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__176; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__177() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__178() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__175; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__176; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__178() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__179() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__57; -x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; -x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__175; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; +x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__59; +x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__176; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__179() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__180() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__178; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__179; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__180() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__181() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__178; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__179; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__181() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__182() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__180; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__181; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__182() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__183() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__179; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__181; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__180; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__182; 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___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__183() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__184() { _start: { lean_object* x_1; @@ -3713,19 +3722,19 @@ x_1 = lean_mk_string_from_bytes("doUnless", 8); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__184() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__185() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__13; -x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__33; -x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__183; +x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__34; +x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__184; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__185() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__186() { _start: { lean_object* x_1; @@ -3733,7 +3742,7 @@ x_1 = lean_mk_string_from_bytes("unless", 6); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__186() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__187() { _start: { lean_object* x_1; @@ -3741,16 +3750,16 @@ x_1 = lean_mk_string_from_bytes("e.isAppOfArity", 14); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__187() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__188() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__186; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__187; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__188() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__189() { _start: { lean_object* x_1; @@ -3758,17 +3767,17 @@ x_1 = lean_mk_string_from_bytes("isAppOfArity", 12); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__189() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__190() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__37; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__188; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__38; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__189; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__190() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__191() { _start: { lean_object* x_1; @@ -3776,19 +3785,19 @@ x_1 = lean_mk_string_from_bytes("dotIdent", 8); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__191() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__192() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__13; -x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__33; -x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__190; +x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__34; +x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__191; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__192() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__193() { _start: { lean_object* x_1; @@ -3796,7 +3805,7 @@ x_1 = lean_mk_string_from_bytes(".", 1); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__193() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__194() { _start: { lean_object* x_1; @@ -3804,26 +3813,26 @@ x_1 = lean_mk_string_from_bytes("continue", 8); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__194() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__195() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__193; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__194; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__195() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__196() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__193; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__194; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__196() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__197() { _start: { lean_object* x_1; @@ -3831,16 +3840,16 @@ x_1 = lean_mk_string_from_bytes("e.appFn!.appArg!", 16); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__197() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__198() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__196; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__197; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__198() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__199() { _start: { lean_object* x_1; @@ -3848,7 +3857,7 @@ x_1 = lean_mk_string_from_bytes("appFn!", 6); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__199() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__200() { _start: { lean_object* x_1; @@ -3856,18 +3865,18 @@ x_1 = lean_mk_string_from_bytes("appArg!", 7); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__200() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__201() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__37; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__198; -x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__199; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__38; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__199; +x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__200; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__201() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__202() { _start: { lean_object* x_1; @@ -3875,26 +3884,26 @@ x_1 = lean_mk_string_from_bytes("m", 1); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__202() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__203() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__201; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__202; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__203() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__204() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__201; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__202; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__204() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__205() { _start: { lean_object* x_1; @@ -3902,26 +3911,26 @@ x_1 = lean_mk_string_from_bytes("e.appArg!", 9); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__205() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__206() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__204; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__205; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__206() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__207() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__37; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__199; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__38; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__200; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__207() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__208() { _start: { lean_object* x_1; @@ -3929,17 +3938,17 @@ x_1 = lean_mk_string_from_bytes("term_<|_", 8); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__208() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__209() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__207; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__208; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__209() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__210() { _start: { lean_object* x_1; @@ -3947,26 +3956,26 @@ x_1 = lean_mk_string_from_bytes("done", 4); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__210() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__211() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__209; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__210; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__211() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__212() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__209; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__210; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__212() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__213() { _start: { lean_object* x_1; @@ -3974,7 +3983,7 @@ x_1 = lean_mk_string_from_bytes("<|", 2); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__213() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__214() { _start: { lean_object* x_1; @@ -3982,26 +3991,26 @@ x_1 = lean_mk_string_from_bytes("toExpr", 6); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__214() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__215() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__213; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__214; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__215() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__216() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__213; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__214; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__216() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__217() { _start: { lean_object* x_1; @@ -4009,42 +4018,42 @@ x_1 = lean_mk_string_from_bytes("ToExpr", 6); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__217() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__218() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__216; -x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__213; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__217; +x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__214; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__218() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__219() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__217; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__218; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__219() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__220() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__218; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__219; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__220() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__221() { _start: { lean_object* x_1; @@ -4052,26 +4061,26 @@ x_1 = lean_mk_string_from_bytes("reduceBinPred", 13); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__221() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__222() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__220; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__221; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__222() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__223() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__220; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__221; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__223() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__224() { _start: { lean_object* x_1; @@ -4079,72 +4088,72 @@ x_1 = lean_mk_string_from_bytes("Bool", 4); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__224() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__225() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__223; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__224; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__225() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__226() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__223; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__224; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__226() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__227() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__225; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__226; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__227() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__228() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__225; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__226; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__228() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__229() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__227; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__228; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__229() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__230() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__226; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__228; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__227; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__229; 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___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__230() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__231() { _start: { lean_object* x_1; @@ -4152,84 +4161,84 @@ x_1 = lean_mk_string_from_bytes("Step", 4); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__231() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__232() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__230; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__231; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__232() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__233() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__230; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__231; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__233() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__234() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__57; -x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; -x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__230; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; +x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__59; +x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__231; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__234() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__235() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__233; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__234; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__235() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__236() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__233; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__234; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__236() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__237() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__235; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__236; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__237() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__238() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__234; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__236; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__235; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__237; 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___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__238() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__239() { _start: { lean_object* x_1; @@ -4237,84 +4246,84 @@ x_1 = lean_mk_string_from_bytes("evalPropStep", 12); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__239() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__240() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__238; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__239; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__240() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__241() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__238; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__239; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__241() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__242() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__57; -x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; -x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__238; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; +x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__59; +x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__239; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__242() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__243() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__241; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__242; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__243() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__244() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__241; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__242; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__244() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__245() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__243; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__244; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__245() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__246() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__242; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__244; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__243; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__245; 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___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__246() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__247() { _start: { lean_object* x_1; @@ -4322,26 +4331,26 @@ x_1 = lean_mk_string_from_bytes("reduceBoolPred", 14); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__247() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__248() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__246; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__247; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__248() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__249() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__246; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__247; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__249() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__250() { _start: { lean_object* x_1; @@ -4349,18 +4358,18 @@ x_1 = lean_mk_string_from_bytes("command__Builtin_dsimproc__[_]_(_):=_", 37); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__250() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__251() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__13; -x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__249; +x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__250; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__251() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__252() { _start: { lean_object* x_1; @@ -4368,7 +4377,7 @@ x_1 = lean_mk_string_from_bytes("builtin_dsimproc", 16); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__252() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__253() { _start: { lean_object* x_1; @@ -4376,7 +4385,7 @@ x_1 = lean_mk_string_from_bytes("[", 1); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__253() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__254() { _start: { lean_object* x_1; @@ -4384,62 +4393,62 @@ x_1 = lean_mk_string_from_bytes("simp", 4); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__254() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__255() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__253; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__254; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__255() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__256() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__253; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__254; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__256() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__257() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__57; -x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; -x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__253; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; +x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__59; +x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__254; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__257() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__258() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__256; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__257; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__258() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__259() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__257; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__258; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__259() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__260() { _start: { lean_object* x_1; @@ -4447,26 +4456,26 @@ x_1 = lean_mk_string_from_bytes("seval", 5); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__260() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__261() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__259; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__260; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__261() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__262() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__259; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__260; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__262() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__263() { _start: { lean_object* x_1; @@ -4474,26 +4483,26 @@ x_1 = lean_mk_string_from_bytes("reduceAdd", 9); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__263() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__264() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__262; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__263; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__264() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__265() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__263; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__264; x_2 = lean_mk_syntax_ident(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__265() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__266() { _start: { lean_object* x_1; @@ -4501,19 +4510,19 @@ x_1 = lean_mk_string_from_bytes("typeAscription", 14); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__266() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__267() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__13; -x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__33; -x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__265; +x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__34; +x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__266; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__267() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__268() { _start: { lean_object* x_1; @@ -4521,17 +4530,17 @@ x_1 = lean_mk_string_from_bytes("term_+_", 7); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__268() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__269() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__267; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__268; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__269() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__270() { _start: { lean_object* x_1; @@ -4539,7 +4548,7 @@ x_1 = lean_mk_string_from_bytes("+", 1); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__270() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__271() { _start: { lean_object* x_1; @@ -4547,19 +4556,19 @@ x_1 = lean_mk_string_from_bytes("doubleQuotedName", 16); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__271() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__272() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__13; -x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__33; -x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__270; +x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__34; +x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__271; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__272() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__273() { _start: { lean_object* x_1; @@ -4567,7 +4576,7 @@ x_1 = lean_mk_string_from_bytes("`", 1); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__273() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__274() { _start: { lean_object* x_1; @@ -4575,16 +4584,16 @@ x_1 = lean_mk_string_from_bytes("HAdd.hAdd", 9); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__274() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__275() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__273; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__274; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__275() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__276() { _start: { lean_object* x_1; @@ -4592,7 +4601,7 @@ x_1 = lean_mk_string_from_bytes("HAdd", 4); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__276() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__277() { _start: { lean_object* x_1; @@ -4600,41 +4609,41 @@ x_1 = lean_mk_string_from_bytes("hAdd", 4); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__277() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__278() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__275; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__276; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__276; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__277; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__278() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__279() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__277; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__278; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__279() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__280() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__278; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__279; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__280() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__281() { _start: { lean_object* x_1; @@ -4642,17 +4651,17 @@ x_1 = lean_mk_string_from_bytes("num", 3); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__281() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__282() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__280; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__281; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__282() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__283() { _start: { lean_object* x_1; @@ -4660,7 +4669,7 @@ x_1 = lean_mk_string_from_bytes("6", 1); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__283() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__284() { _start: { lean_object* x_1; @@ -4668,19 +4677,19 @@ x_1 = lean_mk_string_from_bytes("cdot", 4); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__284() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__285() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__13; -x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__33; -x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__283; +x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__34; +x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__284; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__285() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__286() { _start: { lean_object* x_1; @@ -4688,7 +4697,7 @@ x_1 = lean_mk_string_from_bytes("·", 2); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__286() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__287() { _start: { lean_object* x_1; lean_object* x_2; @@ -4697,7 +4706,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__287() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__288() { _start: { lean_object* x_1; @@ -4705,26 +4714,26 @@ x_1 = lean_mk_string_from_bytes("reduceMul", 9); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__288() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__289() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__287; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__288; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__289() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__290() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__288; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__289; x_2 = lean_mk_syntax_ident(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__290() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__291() { _start: { lean_object* x_1; @@ -4732,17 +4741,17 @@ x_1 = lean_mk_string_from_bytes("term_*_", 7); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__291() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__292() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__290; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__291; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__292() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__293() { _start: { lean_object* x_1; @@ -4750,7 +4759,7 @@ x_1 = lean_mk_string_from_bytes("*", 1); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__293() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__294() { _start: { lean_object* x_1; @@ -4758,16 +4767,16 @@ x_1 = lean_mk_string_from_bytes("HMul.hMul", 9); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__294() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__295() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__293; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__294; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__295() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__296() { _start: { lean_object* x_1; @@ -4775,7 +4784,7 @@ x_1 = lean_mk_string_from_bytes("HMul", 4); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__296() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__297() { _start: { lean_object* x_1; @@ -4783,41 +4792,41 @@ x_1 = lean_mk_string_from_bytes("hMul", 4); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__297() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__298() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__295; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__296; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__296; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__297; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__298() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__299() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__297; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__298; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__299() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__300() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__298; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__299; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__300() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__301() { _start: { lean_object* x_1; @@ -4825,26 +4834,26 @@ x_1 = lean_mk_string_from_bytes("reduceSub", 9); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__301() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__302() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__300; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__301; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__302() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__303() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__301; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__302; x_2 = lean_mk_syntax_ident(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__303() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__304() { _start: { lean_object* x_1; @@ -4852,17 +4861,17 @@ x_1 = lean_mk_string_from_bytes("term_-_", 7); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__304() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__305() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__303; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__304; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__305() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__306() { _start: { lean_object* x_1; @@ -4870,7 +4879,7 @@ x_1 = lean_mk_string_from_bytes("-", 1); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__306() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__307() { _start: { lean_object* x_1; @@ -4878,16 +4887,16 @@ x_1 = lean_mk_string_from_bytes("HSub.hSub", 9); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__307() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__308() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__306; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__307; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__308() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__309() { _start: { lean_object* x_1; @@ -4895,7 +4904,7 @@ x_1 = lean_mk_string_from_bytes("HSub", 4); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__309() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__310() { _start: { lean_object* x_1; @@ -4903,41 +4912,41 @@ x_1 = lean_mk_string_from_bytes("hSub", 4); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__310() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__311() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__308; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__309; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__309; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__310; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__311() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__312() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__310; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__311; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__312() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__313() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__311; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__312; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__313() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__314() { _start: { lean_object* x_1; @@ -4945,26 +4954,26 @@ x_1 = lean_mk_string_from_bytes("reduceDiv", 9); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__314() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__315() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__313; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__314; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__315() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__316() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__314; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__315; x_2 = lean_mk_syntax_ident(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__316() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__317() { _start: { lean_object* x_1; @@ -4972,17 +4981,17 @@ x_1 = lean_mk_string_from_bytes("term_/_", 7); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__317() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__318() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__316; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__317; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__318() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__319() { _start: { lean_object* x_1; @@ -4990,7 +4999,7 @@ x_1 = lean_mk_string_from_bytes("/", 1); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__319() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__320() { _start: { lean_object* x_1; @@ -4998,16 +5007,16 @@ x_1 = lean_mk_string_from_bytes("HDiv.hDiv", 9); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__320() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__321() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__319; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__320; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__321() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__322() { _start: { lean_object* x_1; @@ -5015,7 +5024,7 @@ x_1 = lean_mk_string_from_bytes("HDiv", 4); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__322() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__323() { _start: { lean_object* x_1; @@ -5023,41 +5032,41 @@ x_1 = lean_mk_string_from_bytes("hDiv", 4); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__323() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__324() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__321; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__322; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__322; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__323; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__324() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__325() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__323; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__324; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__325() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__326() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__324; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__325; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__326() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__327() { _start: { lean_object* x_1; @@ -5065,26 +5074,26 @@ x_1 = lean_mk_string_from_bytes("reduceMod", 9); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__327() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__328() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__326; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__327; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__328() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__329() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__327; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__328; x_2 = lean_mk_syntax_ident(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__329() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__330() { _start: { lean_object* x_1; @@ -5092,17 +5101,17 @@ x_1 = lean_mk_string_from_bytes("term_%_", 7); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__330() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__331() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__329; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__330; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__331() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__332() { _start: { lean_object* x_1; @@ -5110,7 +5119,7 @@ x_1 = lean_mk_string_from_bytes("%", 1); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__332() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__333() { _start: { lean_object* x_1; @@ -5118,16 +5127,16 @@ x_1 = lean_mk_string_from_bytes("HMod.hMod", 9); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__333() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__334() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__332; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__333; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__334() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__335() { _start: { lean_object* x_1; @@ -5135,7 +5144,7 @@ x_1 = lean_mk_string_from_bytes("HMod", 4); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__335() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__336() { _start: { lean_object* x_1; @@ -5143,41 +5152,41 @@ x_1 = lean_mk_string_from_bytes("hMod", 4); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__336() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__337() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__334; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__335; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__335; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__336; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__337() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__338() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__336; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__337; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__338() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__339() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__337; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__338; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__339() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__340() { _start: { lean_object* x_1; @@ -5185,18 +5194,18 @@ x_1 = lean_mk_string_from_bytes("command__Builtin_simproc__[_]_(_):=_", 36); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__340() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__341() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__13; -x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__339; +x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__340; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__341() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__342() { _start: { lean_object* x_1; @@ -5204,7 +5213,7 @@ x_1 = lean_mk_string_from_bytes("builtin_simproc", 15); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__342() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__343() { _start: { lean_object* x_1; @@ -5212,26 +5221,26 @@ x_1 = lean_mk_string_from_bytes("reduceLT", 8); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__343() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__344() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__342; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__343; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__344() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__345() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__343; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__344; x_2 = lean_mk_syntax_ident(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__345() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__346() { _start: { lean_object* x_1; @@ -5239,17 +5248,17 @@ x_1 = lean_mk_string_from_bytes("term_<_", 7); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__346() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__347() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__345; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__346; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__347() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__348() { _start: { lean_object* x_1; @@ -5257,7 +5266,7 @@ x_1 = lean_mk_string_from_bytes("<", 1); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__348() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__349() { _start: { lean_object* x_1; @@ -5265,16 +5274,16 @@ x_1 = lean_mk_string_from_bytes("LT.lt", 5); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__349() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__350() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__348; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__349; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__350() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__351() { _start: { lean_object* x_1; @@ -5282,7 +5291,7 @@ x_1 = lean_mk_string_from_bytes("LT", 2); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__351() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__352() { _start: { lean_object* x_1; @@ -5290,41 +5299,41 @@ x_1 = lean_mk_string_from_bytes("lt", 2); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__352() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__353() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__350; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__351; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__351; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__352; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__353() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__354() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__352; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__353; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__354() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__355() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__353; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__354; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__355() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__356() { _start: { lean_object* x_1; @@ -5332,7 +5341,7 @@ x_1 = lean_mk_string_from_bytes("4", 1); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__356() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__357() { _start: { lean_object* x_1; @@ -5340,26 +5349,26 @@ x_1 = lean_mk_string_from_bytes("reduceLE", 8); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__357() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__358() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__356; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__357; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__358() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__359() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__357; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__358; x_2 = lean_mk_syntax_ident(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__359() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__360() { _start: { lean_object* x_1; @@ -5367,17 +5376,17 @@ x_1 = lean_mk_string_from_bytes("term_≤_", 9); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__360() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__361() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__359; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__360; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__361() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__362() { _start: { lean_object* x_1; @@ -5385,7 +5394,7 @@ x_1 = lean_mk_string_from_bytes("≤", 3); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__362() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__363() { _start: { lean_object* x_1; @@ -5393,16 +5402,16 @@ x_1 = lean_mk_string_from_bytes("LE.le", 5); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__363() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__364() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__362; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__363; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__364() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__365() { _start: { lean_object* x_1; @@ -5410,7 +5419,7 @@ x_1 = lean_mk_string_from_bytes("LE", 2); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__365() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__366() { _start: { lean_object* x_1; @@ -5418,41 +5427,41 @@ x_1 = lean_mk_string_from_bytes("le", 2); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__366() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__367() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__364; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__365; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__365; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__366; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__367() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__368() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__366; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__367; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__368() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__369() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__367; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__368; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__369() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__370() { _start: { lean_object* x_1; @@ -5460,26 +5469,26 @@ x_1 = lean_mk_string_from_bytes("reduceGT", 8); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__370() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__371() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__369; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__370; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__371() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__372() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__370; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__371; x_2 = lean_mk_syntax_ident(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__372() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__373() { _start: { lean_object* x_1; @@ -5487,17 +5496,17 @@ x_1 = lean_mk_string_from_bytes("term_>_", 7); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__373() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__374() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__372; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__373; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__374() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__375() { _start: { lean_object* x_1; @@ -5505,7 +5514,7 @@ x_1 = lean_mk_string_from_bytes(">", 1); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__375() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__376() { _start: { lean_object* x_1; @@ -5513,16 +5522,16 @@ x_1 = lean_mk_string_from_bytes("GT.gt", 5); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__376() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__377() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__375; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__376; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__377() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__378() { _start: { lean_object* x_1; @@ -5530,7 +5539,7 @@ x_1 = lean_mk_string_from_bytes("GT", 2); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__378() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__379() { _start: { lean_object* x_1; @@ -5538,41 +5547,41 @@ x_1 = lean_mk_string_from_bytes("gt", 2); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__379() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__380() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__377; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__378; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__378; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__379; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__380() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__381() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__379; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__380; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__381() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__382() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__380; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__381; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__382() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__383() { _start: { lean_object* x_1; @@ -5580,26 +5589,26 @@ x_1 = lean_mk_string_from_bytes("reduceGE", 8); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__383() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__384() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__382; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__383; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__384() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__385() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__383; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__384; x_2 = lean_mk_syntax_ident(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__385() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__386() { _start: { lean_object* x_1; @@ -5607,17 +5616,17 @@ x_1 = lean_mk_string_from_bytes("term_≥_", 9); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__386() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__387() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__385; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__386; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__387() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__388() { _start: { lean_object* x_1; @@ -5625,7 +5634,7 @@ x_1 = lean_mk_string_from_bytes("≥", 3); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__388() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__389() { _start: { lean_object* x_1; @@ -5633,16 +5642,16 @@ x_1 = lean_mk_string_from_bytes("GE.ge", 5); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__389() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__390() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__388; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__389; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__390() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__391() { _start: { lean_object* x_1; @@ -5650,7 +5659,7 @@ x_1 = lean_mk_string_from_bytes("GE", 2); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__391() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__392() { _start: { lean_object* x_1; @@ -5658,41 +5667,41 @@ x_1 = lean_mk_string_from_bytes("ge", 2); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__392() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__393() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__390; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__391; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__391; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__392; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__393() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__394() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__392; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__393; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__394() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__395() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__393; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__394; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__395() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__396() { _start: { lean_object* x_1; @@ -5700,26 +5709,26 @@ x_1 = lean_mk_string_from_bytes("reduceEq", 8); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__396() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__397() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__395; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__396; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__397() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__398() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__395; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__396; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__398() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__399() { _start: { lean_object* x_1; @@ -5727,17 +5736,17 @@ x_1 = lean_mk_string_from_bytes("term_=_", 7); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__399() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__400() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__398; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__399; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__400() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__401() { _start: { lean_object* x_1; @@ -5745,7 +5754,7 @@ x_1 = lean_mk_string_from_bytes("=", 1); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__401() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__402() { _start: { lean_object* x_1; @@ -5753,72 +5762,72 @@ x_1 = lean_mk_string_from_bytes("Eq", 2); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__402() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__403() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__401; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__402; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__403() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__404() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__401; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__402; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__404() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__405() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__403; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__404; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__405() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__406() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__403; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__404; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__406() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__407() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__405; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__406; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__407() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__408() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__404; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__406; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__405; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__407; 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___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__408() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__409() { _start: { lean_object* x_1; @@ -5826,7 +5835,7 @@ x_1 = lean_mk_string_from_bytes("3", 1); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__409() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__410() { _start: { lean_object* x_1; @@ -5834,26 +5843,26 @@ x_1 = lean_mk_string_from_bytes("reduceNe", 8); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__410() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__411() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__409; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__410; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__411() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__412() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__409; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__410; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__412() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__413() { _start: { lean_object* x_1; @@ -5861,17 +5870,17 @@ x_1 = lean_mk_string_from_bytes("term_≠_", 9); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__413() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__414() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__412; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__413; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__414() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__415() { _start: { lean_object* x_1; @@ -5879,7 +5888,7 @@ x_1 = lean_mk_string_from_bytes("≠", 3); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__415() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__416() { _start: { lean_object* x_1; @@ -5887,72 +5896,72 @@ x_1 = lean_mk_string_from_bytes("Ne", 2); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__416() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__417() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__415; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__416; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__417() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__418() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__415; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__416; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__418() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__419() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__417; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__418; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__419() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__420() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__417; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__418; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__420() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__421() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__419; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__420; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__421() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__422() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__418; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__420; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__419; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__421; 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___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__422() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__423() { _start: { lean_object* x_1; @@ -5960,26 +5969,26 @@ x_1 = lean_mk_string_from_bytes("reduceBEq", 9); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__423() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__424() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__422; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__423; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__424() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__425() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__422; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__423; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__425() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__426() { _start: { lean_object* x_1; @@ -5987,17 +5996,17 @@ x_1 = lean_mk_string_from_bytes("term_==_", 8); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__426() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__427() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__425; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__426; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__427() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__428() { _start: { lean_object* x_1; @@ -6005,7 +6014,7 @@ x_1 = lean_mk_string_from_bytes("==", 2); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__428() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__429() { _start: { lean_object* x_1; @@ -6013,16 +6022,16 @@ x_1 = lean_mk_string_from_bytes("BEq.beq", 7); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__429() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__430() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__428; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__429; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__430() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__431() { _start: { lean_object* x_1; @@ -6030,7 +6039,7 @@ x_1 = lean_mk_string_from_bytes("BEq", 3); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__431() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__432() { _start: { lean_object* x_1; @@ -6038,41 +6047,41 @@ x_1 = lean_mk_string_from_bytes("beq", 3); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__432() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__433() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__430; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__431; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__431; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__432; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__433() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__434() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__432; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__433; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__434() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__435() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__433; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__434; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__435() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__436() { _start: { lean_object* x_1; @@ -6080,26 +6089,26 @@ x_1 = lean_mk_string_from_bytes("reduceBNe", 9); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__436() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__437() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__435; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__436; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__437() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__438() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__435; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__436; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__438() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__439() { _start: { lean_object* x_1; @@ -6107,17 +6116,17 @@ x_1 = lean_mk_string_from_bytes("term_!=_", 8); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__439() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__440() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__438; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__439; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__440() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__441() { _start: { lean_object* x_1; @@ -6125,7 +6134,7 @@ x_1 = lean_mk_string_from_bytes("!=", 2); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__441() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__442() { _start: { lean_object* x_1; @@ -6133,72 +6142,72 @@ x_1 = lean_mk_string_from_bytes("bne", 3); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__442() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__443() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__441; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__442; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__443() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__444() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__441; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__442; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__444() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__445() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__443; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__444; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__445() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__446() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__443; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__444; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__446() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__447() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__445; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__446; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__447() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__448() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__444; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__446; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__445; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__447; 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___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__448() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__449() { _start: { lean_object* x_1; @@ -6206,26 +6215,26 @@ x_1 = lean_mk_string_from_bytes("reduceOfNatCore", 15); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__449() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__450() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__448; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__449; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__450() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__451() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__449; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__450; x_2 = lean_mk_syntax_ident(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__451() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__452() { _start: { lean_object* x_1; @@ -6233,19 +6242,19 @@ x_1 = lean_mk_string_from_bytes("fun", 3); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__452() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__453() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__13; -x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__33; -x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__451; +x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__34; +x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__452; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__453() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__454() { _start: { lean_object* x_1; @@ -6253,19 +6262,19 @@ x_1 = lean_mk_string_from_bytes("basicFun", 8); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__454() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__455() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__13; -x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__33; -x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__453; +x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__34; +x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__454; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__455() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__456() { _start: { lean_object* x_1; @@ -6273,7 +6282,7 @@ x_1 = lean_mk_string_from_bytes("=>", 2); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__456() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__457() { _start: { lean_object* x_1; @@ -6281,7 +6290,7 @@ x_1 = lean_mk_string_from_bytes("2", 1); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__457() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__458() { _start: { lean_object* x_1; @@ -6289,26 +6298,26 @@ x_1 = lean_mk_string_from_bytes("value", 5); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__458() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__459() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__457; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__458; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__459() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__460() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__457; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__458; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__460() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__461() { _start: { lean_object* x_1; @@ -6316,16 +6325,16 @@ x_1 = lean_mk_string_from_bytes("Nat.fromExpr\?", 13); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__461() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__462() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__460; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__461; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__462() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__463() { _start: { lean_object* x_1; @@ -6333,63 +6342,63 @@ x_1 = lean_mk_string_from_bytes("fromExpr\?", 9); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__463() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__464() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__162; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__462; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__163; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__463; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__464() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__465() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__463; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__464; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__465() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__466() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__463; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__464; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__466() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__467() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__465; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__466; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__467() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__468() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__464; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__466; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__465; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__467; 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___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__468() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__469() { _start: { lean_object* x_1; @@ -6397,19 +6406,19 @@ x_1 = lean_mk_string_from_bytes("doLet", 5); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__469() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__470() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__13; -x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__33; -x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__468; +x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__34; +x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__469; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__470() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__471() { _start: { lean_object* x_1; @@ -6417,19 +6426,19 @@ x_1 = lean_mk_string_from_bytes("letDecl", 7); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__471() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__472() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__13; -x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__33; -x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__470; +x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__34; +x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__471; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__472() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__473() { _start: { lean_object* x_1; @@ -6437,19 +6446,19 @@ x_1 = lean_mk_string_from_bytes("letIdDecl", 9); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__473() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__474() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__13; -x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__33; -x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__472; +x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__34; +x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__473; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__474() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__475() { _start: { lean_object* x_1; @@ -6457,26 +6466,26 @@ x_1 = lean_mk_string_from_bytes("reduceOfNat", 11); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__475() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__476() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__474; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__475; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__476() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__477() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__475; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__476; x_2 = lean_mk_syntax_ident(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__477() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__478() { _start: { lean_object* x_1; @@ -6484,7 +6493,7 @@ x_1 = lean_mk_string_from_bytes("1", 1); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__478() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__479() { _start: { lean_object* x_1; @@ -6492,26 +6501,26 @@ x_1 = lean_mk_string_from_bytes("reduceToNat", 11); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__479() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__480() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__478; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__479; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__480() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__481() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__479; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__480; x_2 = lean_mk_syntax_ident(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__481() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__482() { _start: { lean_object* x_1; @@ -6519,26 +6528,26 @@ x_1 = lean_mk_string_from_bytes("v", 1); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__482() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__483() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__481; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__482; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__483() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__484() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__481; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__482; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__484() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__485() { _start: { lean_object* x_1; @@ -6546,19 +6555,19 @@ x_1 = lean_mk_string_from_bytes("docComment", 10); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__485() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__486() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__13; x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__14; -x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__484; +x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__485; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__486() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__487() { _start: { lean_object* x_1; @@ -6566,7 +6575,7 @@ x_1 = lean_mk_string_from_bytes("/--", 3); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__487() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__488() { _start: { lean_object* x_1; @@ -6574,7 +6583,7 @@ x_1 = lean_mk_string_from_bytes("Return `.done` for UInt values. We don't want t return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__488() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__489() { _start: { lean_object* x_1; @@ -6582,26 +6591,26 @@ x_1 = lean_mk_string_from_bytes("isValue", 7); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__489() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__490() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__488; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__489; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__490() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__491() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__488; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__489; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__491() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__492() { _start: { lean_object* x_1; @@ -6609,16 +6618,16 @@ x_1 = lean_mk_string_from_bytes("OfNat.ofNat", 11); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__492() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__493() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__491; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__492; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__493() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__494() { _start: { lean_object* x_1; @@ -6626,41 +6635,41 @@ x_1 = lean_mk_string_from_bytes("OfNat", 5); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__494() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__495() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__493; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__494; x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__495() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__496() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__494; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__495; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__496() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__497() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__495; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__496; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__497() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__498() { _start: { lean_object* x_1; @@ -6668,19 +6677,19 @@ x_1 = lean_mk_string_from_bytes("end", 3); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__498() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__499() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__13; x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__14; -x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__497; +x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__498; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__499() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__500() { _start: { lean_object* x_1; lean_object* x_2; @@ -6689,7 +6698,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__500() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__501() { _start: { lean_object* x_1; @@ -6697,19 +6706,19 @@ x_1 = lean_mk_string_from_bytes("quotedName", 10); return x_1; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__501() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__502() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__12; x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__13; -x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__33; -x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__500; +x_3 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__34; +x_4 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__501; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__502() { +static lean_object* _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__503() { _start: { lean_object* x_1; lean_object* x_2; @@ -6783,22 +6792,22 @@ x_31 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules_ lean_inc_n(x_30, 6); lean_inc(x_21); x_32 = l_Lean_Syntax_node6(x_21, x_31, x_30, x_30, x_30, x_30, x_30, x_30); -x_33 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__22; +x_33 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__24; lean_inc(x_21); x_34 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_34, 0, x_21); lean_ctor_set(x_34, 1, x_33); -x_35 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__36; +x_35 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__37; lean_inc(x_21); x_36 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_36, 0, x_21); lean_ctor_set(x_36, 1, x_35); -x_37 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__39; +x_37 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__40; lean_inc(x_22); lean_inc(x_23); x_38 = l_Lean_addMacroScope(x_23, x_37, x_22); x_39 = lean_box(0); -x_40 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__38; +x_40 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__39; lean_inc(x_21); x_41 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_41, 0, x_21); @@ -6808,17 +6817,17 @@ lean_ctor_set(x_41, 3, x_39); lean_inc(x_41); lean_inc(x_21); x_42 = l_Lean_Syntax_node1(x_21, x_28, x_41); -x_43 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__40; +x_43 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__41; lean_inc(x_21); x_44 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_44, 0, x_21); lean_ctor_set(x_44, 1, x_43); -x_45 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__43; +x_45 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__44; lean_inc(x_22); lean_inc(x_23); x_46 = l_Lean_addMacroScope(x_23, x_45, x_22); -x_47 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__42; -x_48 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__48; +x_47 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__43; +x_48 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__49; lean_inc(x_21); x_49 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_49, 0, x_21); @@ -6828,12 +6837,12 @@ lean_ctor_set(x_49, 3, x_48); lean_inc(x_44); lean_inc(x_21); x_50 = l_Lean_Syntax_node2(x_21, x_28, x_44, x_49); -x_51 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__49; +x_51 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__50; lean_inc(x_21); x_52 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_52, 0, x_21); lean_ctor_set(x_52, 1, x_51); -x_53 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__35; +x_53 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__36; lean_inc(x_52); lean_inc(x_30); lean_inc(x_42); @@ -6843,24 +6852,24 @@ x_54 = l_Lean_Syntax_node5(x_21, x_53, x_36, x_42, x_50, x_30, x_52); lean_inc(x_54); lean_inc(x_21); x_55 = l_Lean_Syntax_node1(x_21, x_28, x_54); -x_56 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__56; +x_56 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__57; lean_inc(x_22); lean_inc(x_23); x_57 = l_Lean_addMacroScope(x_23, x_56, x_22); -x_58 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__55; -x_59 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__61; +x_58 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__56; +x_59 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__62; lean_inc(x_21); x_60 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_60, 0, x_21); lean_ctor_set(x_60, 1, x_58); lean_ctor_set(x_60, 2, x_57); lean_ctor_set(x_60, 3, x_59); -x_61 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__66; +x_61 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__67; lean_inc(x_22); lean_inc(x_23); x_62 = l_Lean_addMacroScope(x_23, x_61, x_22); -x_63 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__65; -x_64 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__73; +x_63 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__66; +x_64 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__74; lean_inc(x_21); x_65 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_65, 0, x_21); @@ -6870,11 +6879,11 @@ lean_ctor_set(x_65, 3, x_64); lean_inc(x_9); lean_inc(x_21); x_66 = l_Lean_Syntax_node1(x_21, x_28, x_9); -x_67 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__53; +x_67 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__54; lean_inc(x_66); lean_inc(x_21); x_68 = l_Lean_Syntax_node2(x_21, x_67, x_65, x_66); -x_69 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__63; +x_69 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__64; lean_inc(x_52); lean_inc(x_36); lean_inc(x_21); @@ -6884,64 +6893,64 @@ x_71 = l_Lean_Syntax_node1(x_21, x_28, x_70); lean_inc(x_60); lean_inc(x_21); x_72 = l_Lean_Syntax_node2(x_21, x_67, x_60, x_71); -x_73 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__51; +x_73 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__52; lean_inc(x_44); lean_inc(x_21); x_74 = l_Lean_Syntax_node2(x_21, x_73, x_44, x_72); lean_inc(x_21); x_75 = l_Lean_Syntax_node1(x_21, x_28, x_74); -x_76 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__32; +x_76 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__33; lean_inc(x_21); x_77 = l_Lean_Syntax_node2(x_21, x_76, x_55, x_75); -x_78 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__76; +x_78 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__77; lean_inc(x_21); x_79 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_79, 0, x_21); lean_ctor_set(x_79, 1, x_78); -x_80 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__77; +x_80 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__78; lean_inc(x_21); x_81 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_81, 0, x_21); lean_ctor_set(x_81, 1, x_80); -x_82 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__85; +x_82 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__86; lean_inc(x_21); x_83 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_83, 0, x_21); lean_ctor_set(x_83, 1, x_82); -x_84 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__90; +x_84 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__91; lean_inc(x_22); lean_inc(x_23); x_85 = l_Lean_addMacroScope(x_23, x_84, x_22); -x_86 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__89; -x_87 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__93; +x_86 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__90; +x_87 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__94; lean_inc(x_21); x_88 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_88, 0, x_21); lean_ctor_set(x_88, 1, x_86); lean_ctor_set(x_88, 2, x_85); lean_ctor_set(x_88, 3, x_87); -x_89 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__98; +x_89 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__99; lean_inc(x_22); lean_inc(x_23); x_90 = l_Lean_addMacroScope(x_23, x_89, x_22); -x_91 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__97; +x_91 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__98; lean_inc(x_21); x_92 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_92, 0, x_21); lean_ctor_set(x_92, 1, x_91); lean_ctor_set(x_92, 2, x_90); lean_ctor_set(x_92, 3, x_39); -x_93 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__99; +x_93 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__100; lean_inc(x_21); x_94 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_94, 0, x_21); lean_ctor_set(x_94, 1, x_93); -x_95 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__102; +x_95 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__103; lean_inc(x_21); x_96 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_96, 0, x_21); lean_ctor_set(x_96, 1, x_95); -x_97 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__101; +x_97 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__102; lean_inc(x_21); x_98 = l_Lean_Syntax_node1(x_21, x_97, x_96); lean_inc(x_98); @@ -6952,7 +6961,7 @@ lean_inc(x_94); lean_inc(x_92); lean_inc(x_21); x_100 = l_Lean_Syntax_node3(x_21, x_28, x_92, x_94, x_99); -x_101 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__95; +x_101 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__96; lean_inc(x_52); lean_inc(x_36); lean_inc(x_21); @@ -6962,17 +6971,17 @@ x_103 = l_Lean_Syntax_node1(x_21, x_28, x_102); lean_inc(x_88); lean_inc(x_21); x_104 = l_Lean_Syntax_node2(x_21, x_67, x_88, x_103); -x_105 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__103; +x_105 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__104; 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___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__108; +x_107 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__109; lean_inc(x_22); lean_inc(x_23); x_108 = l_Lean_addMacroScope(x_23, x_107, x_22); -x_109 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__107; -x_110 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__113; +x_109 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__108; +x_110 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__114; lean_inc(x_21); x_111 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_111, 0, x_21); @@ -6981,22 +6990,22 @@ lean_ctor_set(x_111, 2, x_108); lean_ctor_set(x_111, 3, x_110); lean_inc(x_10); x_112 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_39, x_10); -x_113 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__114; +x_113 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__115; lean_inc(x_21); x_114 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_114, 0, x_21); lean_ctor_set(x_114, 1, x_113); -x_115 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__117; +x_115 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__118; lean_inc(x_21); x_116 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_116, 0, x_21); lean_ctor_set(x_116, 1, x_115); -x_117 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__120; +x_117 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__121; lean_inc(x_22); lean_inc(x_23); x_118 = l_Lean_addMacroScope(x_23, x_117, x_22); -x_119 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__119; -x_120 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__123; +x_119 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__120; +x_120 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__124; lean_inc(x_21); x_121 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_121, 0, x_21); @@ -7005,17 +7014,17 @@ lean_ctor_set(x_121, 2, x_118); lean_ctor_set(x_121, 3, x_120); lean_inc(x_21); x_122 = l_Lean_Syntax_node1(x_21, x_28, x_121); -x_123 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__116; +x_123 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__117; lean_inc(x_116); lean_inc(x_21); x_124 = l_Lean_Syntax_node2(x_21, x_123, x_116, x_122); -x_125 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__82; +x_125 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__83; lean_inc(x_30); lean_inc(x_21); x_126 = l_Lean_Syntax_node2(x_21, x_125, x_124, x_30); lean_inc(x_21); x_127 = l_Lean_Syntax_node1(x_21, x_28, x_126); -x_128 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__80; +x_128 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__81; lean_inc(x_21); x_129 = l_Lean_Syntax_node1(x_21, x_128, x_127); lean_inc(x_114); @@ -7038,47 +7047,47 @@ x_135 = l_Lean_Syntax_node2(x_21, x_123, x_116, x_134); lean_inc(x_30); lean_inc(x_21); x_136 = l_Lean_Syntax_node2(x_21, x_125, x_135, x_30); -x_137 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__126; +x_137 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__127; lean_inc_n(x_30, 2); lean_inc(x_21); x_138 = l_Lean_Syntax_node2(x_21, x_137, x_30, x_30); -x_139 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__129; +x_139 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__130; lean_inc(x_21); x_140 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_140, 0, x_21); lean_ctor_set(x_140, 1, x_139); -x_141 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__133; +x_141 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__134; lean_inc(x_30); lean_inc(x_21); x_142 = l_Lean_Syntax_node1(x_21, x_141, x_30); -x_143 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__139; +x_143 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__140; lean_inc(x_22); lean_inc(x_23); x_144 = l_Lean_addMacroScope(x_23, x_143, x_22); -x_145 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__138; -x_146 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__143; +x_145 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__139; +x_146 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__144; lean_inc(x_21); x_147 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_147, 0, x_21); lean_ctor_set(x_147, 1, x_145); lean_ctor_set(x_147, 2, x_144); lean_ctor_set(x_147, 3, x_146); -x_148 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__136; +x_148 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__137; lean_inc(x_30); lean_inc(x_21); x_149 = l_Lean_Syntax_node2(x_21, x_148, x_147, x_30); -x_150 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__131; +x_150 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__132; lean_inc(x_142); lean_inc(x_21); x_151 = l_Lean_Syntax_node2(x_21, x_150, x_142, x_149); lean_inc(x_21); x_152 = l_Lean_Syntax_node1(x_21, x_28, x_151); -x_153 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__144; +x_153 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__145; lean_inc(x_21); x_154 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_154, 0, x_21); lean_ctor_set(x_154, 1, x_153); -x_155 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__128; +x_155 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__129; lean_inc(x_154); lean_inc(x_21); x_156 = l_Lean_Syntax_node3(x_21, x_155, x_140, x_152, x_154); @@ -7087,27 +7096,27 @@ x_157 = l_Lean_Syntax_node1(x_21, x_28, x_156); lean_inc_n(x_30, 5); lean_inc(x_21); x_158 = l_Lean_Syntax_node6(x_21, x_31, x_30, x_157, x_30, x_30, x_30, x_30); -x_159 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__147; +x_159 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__148; lean_inc(x_22); lean_inc(x_23); x_160 = l_Lean_addMacroScope(x_23, x_159, x_22); -x_161 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__146; +x_161 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__147; lean_inc(x_21); x_162 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_162, 0, x_21); lean_ctor_set(x_162, 1, x_161); lean_ctor_set(x_162, 2, x_160); lean_ctor_set(x_162, 3, x_39); -x_163 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__25; +x_163 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__26; lean_inc(x_30); lean_inc(x_162); lean_inc(x_21); x_164 = l_Lean_Syntax_node2(x_21, x_163, x_162, x_30); -x_165 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__150; +x_165 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__151; lean_inc(x_22); lean_inc(x_23); x_166 = l_Lean_addMacroScope(x_23, x_165, x_22); -x_167 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__149; +x_167 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__150; lean_inc(x_21); x_168 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_168, 0, x_21); @@ -7117,12 +7126,12 @@ lean_ctor_set(x_168, 3, x_39); lean_inc(x_168); lean_inc(x_21); x_169 = l_Lean_Syntax_node1(x_21, x_28, x_168); -x_170 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__153; +x_170 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__154; lean_inc(x_22); lean_inc(x_23); x_171 = l_Lean_addMacroScope(x_23, x_170, x_22); -x_172 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__152; -x_173 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__158; +x_172 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__153; +x_173 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__159; lean_inc(x_21); x_174 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_174, 0, x_21); @@ -7137,11 +7146,11 @@ lean_inc(x_30); lean_inc(x_36); lean_inc(x_21); x_176 = l_Lean_Syntax_node5(x_21, x_53, x_36, x_169, x_175, x_30, x_52); -x_177 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__161; +x_177 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__162; lean_inc(x_22); lean_inc(x_23); x_178 = l_Lean_addMacroScope(x_23, x_177, x_22); -x_179 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__160; +x_179 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__161; lean_inc(x_21); x_180 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_180, 0, x_21); @@ -7151,12 +7160,12 @@ lean_ctor_set(x_180, 3, x_39); lean_inc(x_180); lean_inc(x_21); x_181 = l_Lean_Syntax_node1(x_21, x_28, x_180); -x_182 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__164; +x_182 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__165; lean_inc(x_22); lean_inc(x_23); x_183 = l_Lean_addMacroScope(x_23, x_182, x_22); -x_184 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__163; -x_185 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__168; +x_184 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__164; +x_185 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__169; lean_inc(x_21); x_186 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_186, 0, x_21); @@ -7171,11 +7180,11 @@ lean_inc(x_30); lean_inc(x_36); lean_inc(x_21); x_188 = l_Lean_Syntax_node5(x_21, x_53, x_36, x_181, x_187, x_30, x_52); -x_189 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__171; +x_189 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__172; lean_inc(x_22); lean_inc(x_23); x_190 = l_Lean_addMacroScope(x_23, x_189, x_22); -x_191 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__170; +x_191 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__171; lean_inc(x_21); x_192 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_192, 0, x_21); @@ -7185,12 +7194,12 @@ lean_ctor_set(x_192, 3, x_39); lean_inc(x_192); lean_inc(x_21); x_193 = l_Lean_Syntax_node1(x_21, x_28, x_192); -x_194 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__174; +x_194 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__175; lean_inc(x_21); x_195 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_195, 0, x_21); lean_ctor_set(x_195, 1, x_194); -x_196 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__173; +x_196 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__174; lean_inc(x_195); lean_inc_n(x_9, 2); lean_inc(x_21); @@ -7213,12 +7222,12 @@ lean_inc(x_188); lean_inc(x_176); lean_inc(x_21); x_201 = l_Lean_Syntax_node4(x_21, x_28, x_176, x_188, x_200, x_54); -x_202 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__177; +x_202 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__178; lean_inc(x_22); lean_inc(x_23); x_203 = l_Lean_addMacroScope(x_23, x_202, x_22); -x_204 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__176; -x_205 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__182; +x_204 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__177; +x_205 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__183; lean_inc(x_21); x_206 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_206, 0, x_21); @@ -7238,16 +7247,16 @@ x_210 = l_Lean_Syntax_node1(x_21, x_28, x_209); lean_inc(x_210); lean_inc(x_21); x_211 = l_Lean_Syntax_node2(x_21, x_76, x_201, x_210); -x_212 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__185; +x_212 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__186; lean_inc(x_21); x_213 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_213, 0, x_21); lean_ctor_set(x_213, 1, x_212); -x_214 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__189; +x_214 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__190; lean_inc(x_22); lean_inc(x_23); x_215 = l_Lean_addMacroScope(x_23, x_214, x_22); -x_216 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__187; +x_216 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__188; lean_inc(x_21); x_217 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_217, 0, x_21); @@ -7259,23 +7268,23 @@ x_218 = l_Lean_Syntax_node2(x_21, x_28, x_168, x_180); lean_inc(x_217); lean_inc(x_21); x_219 = l_Lean_Syntax_node2(x_21, x_67, x_217, x_218); -x_220 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__192; +x_220 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__193; lean_inc(x_21); x_221 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_221, 0, x_21); lean_ctor_set(x_221, 1, x_220); -x_222 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__195; +x_222 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__196; lean_inc(x_22); lean_inc(x_23); x_223 = l_Lean_addMacroScope(x_23, x_222, x_22); -x_224 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__194; +x_224 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__195; lean_inc(x_21); x_225 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_225, 0, x_21); lean_ctor_set(x_225, 1, x_224); lean_ctor_set(x_225, 2, x_223); lean_ctor_set(x_225, 3, x_39); -x_226 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__191; +x_226 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__192; lean_inc(x_221); lean_inc(x_21); x_227 = l_Lean_Syntax_node2(x_21, x_226, x_221, x_225); @@ -7291,7 +7300,7 @@ lean_inc(x_21); x_231 = l_Lean_Syntax_node1(x_21, x_28, x_230); lean_inc(x_21); x_232 = l_Lean_Syntax_node1(x_21, x_128, x_231); -x_233 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__184; +x_233 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__185; lean_inc(x_232); lean_inc(x_81); lean_inc(x_213); @@ -7304,11 +7313,11 @@ lean_inc(x_132); lean_inc(x_88); lean_inc(x_21); x_236 = l_Lean_Syntax_node2(x_21, x_67, x_88, x_132); -x_237 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__200; +x_237 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__201; lean_inc(x_22); lean_inc(x_23); x_238 = l_Lean_addMacroScope(x_23, x_237, x_22); -x_239 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__197; +x_239 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__198; lean_inc(x_21); x_240 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_240, 0, x_21); @@ -7325,18 +7334,18 @@ lean_inc(x_52); lean_inc(x_36); lean_inc(x_21); x_244 = l_Lean_Syntax_node3(x_21, x_69, x_36, x_243, x_52); -x_245 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__105; +x_245 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__106; lean_inc(x_21); x_246 = l_Lean_Syntax_node1(x_21, x_245, x_244); lean_inc(x_232); lean_inc(x_21); x_247 = l_Lean_Syntax_node2(x_21, x_28, x_114, x_232); -x_248 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__87; +x_248 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__88; lean_inc(x_247); lean_inc(x_106); lean_inc(x_21); x_249 = l_Lean_Syntax_node4(x_21, x_248, x_236, x_106, x_246, x_247); -x_250 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__84; +x_250 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__85; lean_inc(x_30); lean_inc(x_83); lean_inc(x_21); @@ -7344,11 +7353,11 @@ x_251 = l_Lean_Syntax_node3(x_21, x_250, x_83, x_30, x_249); lean_inc(x_30); lean_inc(x_21); x_252 = l_Lean_Syntax_node2(x_21, x_125, x_251, x_30); -x_253 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__203; +x_253 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__204; lean_inc(x_22); lean_inc(x_23); x_254 = l_Lean_addMacroScope(x_23, x_253, x_22); -x_255 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__202; +x_255 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__203; lean_inc(x_21); x_256 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_256, 0, x_21); @@ -7361,11 +7370,11 @@ x_257 = l_Lean_Syntax_node1(x_21, x_28, x_256); lean_inc(x_88); lean_inc(x_21); x_258 = l_Lean_Syntax_node2(x_21, x_67, x_88, x_257); -x_259 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__206; +x_259 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__207; lean_inc(x_22); lean_inc(x_23); x_260 = l_Lean_addMacroScope(x_23, x_259, x_22); -x_261 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__205; +x_261 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__206; lean_inc(x_21); x_262 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_262, 0, x_21); @@ -7395,11 +7404,11 @@ x_268 = l_Lean_Syntax_node3(x_21, x_250, x_83, x_30, x_267); lean_inc(x_30); lean_inc(x_21); x_269 = l_Lean_Syntax_node2(x_21, x_125, x_268, x_30); -x_270 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__211; +x_270 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__212; lean_inc(x_22); lean_inc(x_23); x_271 = l_Lean_addMacroScope(x_23, x_270, x_22); -x_272 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__210; +x_272 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__211; lean_inc(x_21); x_273 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_273, 0, x_21); @@ -7409,17 +7418,17 @@ lean_ctor_set(x_273, 3, x_39); lean_inc(x_221); lean_inc(x_21); x_274 = l_Lean_Syntax_node2(x_21, x_226, x_221, x_273); -x_275 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__212; +x_275 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__213; lean_inc(x_21); x_276 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_276, 0, x_21); lean_ctor_set(x_276, 1, x_275); -x_277 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__215; +x_277 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__216; lean_inc(x_22); lean_inc(x_23); x_278 = l_Lean_addMacroScope(x_23, x_277, x_22); -x_279 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__214; -x_280 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__219; +x_279 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__215; +x_280 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__220; lean_inc(x_21); x_281 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_281, 0, x_21); @@ -7441,7 +7450,7 @@ x_285 = l_Lean_Syntax_node1(x_21, x_28, x_284); lean_inc(x_281); lean_inc(x_21); x_286 = l_Lean_Syntax_node2(x_21, x_67, x_281, x_285); -x_287 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__208; +x_287 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__209; lean_inc(x_276); lean_inc(x_274); lean_inc(x_21); @@ -7461,11 +7470,11 @@ lean_inc(x_21); x_292 = l_Lean_Syntax_node4(x_21, x_28, x_235, x_252, x_269, x_291); lean_inc(x_21); x_293 = l_Lean_Syntax_node1(x_21, x_128, x_292); -x_294 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__78; +x_294 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__79; lean_inc(x_81); lean_inc(x_21); x_295 = l_Lean_Syntax_node2(x_21, x_294, x_81, x_293); -x_296 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__75; +x_296 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__76; lean_inc(x_30); lean_inc(x_138); lean_inc(x_79); @@ -7481,11 +7490,11 @@ x_300 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules lean_inc(x_158); lean_inc(x_21); x_301 = l_Lean_Syntax_node2(x_21, x_300, x_158, x_299); -x_302 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__222; +x_302 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__223; lean_inc(x_22); lean_inc(x_23); x_303 = l_Lean_addMacroScope(x_23, x_302, x_22); -x_304 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__221; +x_304 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__222; lean_inc(x_21); x_305 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_305, 0, x_21); @@ -7496,12 +7505,12 @@ lean_inc(x_30); lean_inc(x_305); lean_inc(x_21); x_306 = l_Lean_Syntax_node2(x_21, x_163, x_305, x_30); -x_307 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__225; +x_307 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__226; lean_inc(x_22); lean_inc(x_23); x_308 = l_Lean_addMacroScope(x_23, x_307, x_22); -x_309 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__224; -x_310 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__229; +x_309 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__225; +x_310 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__230; lean_inc(x_21); x_311 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_311, 0, x_21); @@ -7524,12 +7533,12 @@ lean_inc(x_21); x_315 = l_Lean_Syntax_node5(x_21, x_53, x_36, x_193, x_314, x_30, x_52); lean_inc(x_21); x_316 = l_Lean_Syntax_node4(x_21, x_28, x_176, x_188, x_315, x_54); -x_317 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__232; +x_317 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__233; lean_inc(x_22); lean_inc(x_23); x_318 = l_Lean_addMacroScope(x_23, x_317, x_22); -x_319 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__231; -x_320 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__237; +x_319 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__232; +x_320 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__238; lean_inc(x_21); x_321 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_321, 0, x_21); @@ -7548,12 +7557,12 @@ x_325 = l_Lean_Syntax_node1(x_21, x_28, x_324); lean_inc(x_316); lean_inc(x_21); x_326 = l_Lean_Syntax_node2(x_21, x_76, x_316, x_325); -x_327 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__240; +x_327 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__241; lean_inc(x_22); lean_inc(x_23); x_328 = l_Lean_addMacroScope(x_23, x_327, x_22); -x_329 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__239; -x_330 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__245; +x_329 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__240; +x_330 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__246; lean_inc(x_21); x_331 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_331, 0, x_21); @@ -7589,11 +7598,11 @@ x_340 = l_Lean_Syntax_node5(x_21, x_298, x_34, x_306, x_326, x_339, x_30); lean_inc(x_158); lean_inc(x_21); x_341 = l_Lean_Syntax_node2(x_21, x_300, x_158, x_340); -x_342 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__248; +x_342 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__249; lean_inc(x_22); lean_inc(x_23); x_343 = l_Lean_addMacroScope(x_23, x_342, x_22); -x_344 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__247; +x_344 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__248; lean_inc(x_21); x_345 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_345, 0, x_21); @@ -7612,33 +7621,33 @@ lean_inc(x_21); x_348 = l_Lean_Syntax_node5(x_21, x_298, x_34, x_346, x_347, x_297, x_30); lean_inc(x_21); x_349 = l_Lean_Syntax_node2(x_21, x_300, x_158, x_348); -x_350 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__251; +x_350 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__252; lean_inc(x_21); x_351 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_351, 0, x_21); lean_ctor_set(x_351, 1, x_350); -x_352 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__252; +x_352 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__253; lean_inc(x_21); x_353 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_353, 0, x_21); lean_ctor_set(x_353, 1, x_352); -x_354 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__255; +x_354 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__256; lean_inc(x_22); lean_inc(x_23); x_355 = l_Lean_addMacroScope(x_23, x_354, x_22); -x_356 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__254; -x_357 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__258; +x_356 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__255; +x_357 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__259; lean_inc(x_21); x_358 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_358, 0, x_21); lean_ctor_set(x_358, 1, x_356); lean_ctor_set(x_358, 2, x_355); lean_ctor_set(x_358, 3, x_357); -x_359 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__261; +x_359 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__262; lean_inc(x_22); lean_inc(x_23); x_360 = l_Lean_addMacroScope(x_23, x_359, x_22); -x_361 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__260; +x_361 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__261; lean_inc(x_21); x_362 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_362, 0, x_21); @@ -7652,58 +7661,58 @@ lean_inc(x_154); lean_inc(x_353); lean_inc(x_21); x_364 = l_Lean_Syntax_node3(x_21, x_28, x_353, x_363, x_154); -x_365 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__269; +x_365 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__270; lean_inc(x_21); x_366 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_366, 0, x_21); lean_ctor_set(x_366, 1, x_365); -x_367 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__268; +x_367 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__269; lean_inc(x_366); lean_inc_n(x_98, 2); lean_inc(x_21); x_368 = l_Lean_Syntax_node3(x_21, x_367, x_98, x_366, x_98); -x_369 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__266; +x_369 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__267; lean_inc(x_52); lean_inc(x_66); lean_inc(x_44); lean_inc(x_36); lean_inc(x_21); x_370 = l_Lean_Syntax_node5(x_21, x_369, x_36, x_368, x_44, x_66, x_52); -x_371 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__272; +x_371 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__273; lean_inc(x_21); x_372 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_372, 0, x_21); lean_ctor_set(x_372, 1, x_371); -x_373 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__277; +x_373 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__278; lean_inc(x_22); lean_inc(x_23); x_374 = l_Lean_addMacroScope(x_23, x_373, x_22); -x_375 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__274; -x_376 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__279; +x_375 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__275; +x_376 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__280; lean_inc(x_21); x_377 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_377, 0, x_21); lean_ctor_set(x_377, 1, x_375); lean_ctor_set(x_377, 2, x_374); lean_ctor_set(x_377, 3, x_376); -x_378 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__271; +x_378 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__272; lean_inc_n(x_372, 2); lean_inc(x_21); x_379 = l_Lean_Syntax_node3(x_21, x_378, x_372, x_372, x_377); -x_380 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__282; +x_380 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__283; lean_inc(x_21); x_381 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_381, 0, x_21); lean_ctor_set(x_381, 1, x_380); -x_382 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__281; +x_382 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__282; lean_inc(x_21); x_383 = l_Lean_Syntax_node1(x_21, x_382, x_381); -x_384 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__285; +x_384 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__286; lean_inc(x_21); x_385 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_385, 0, x_21); lean_ctor_set(x_385, 1, x_384); -x_386 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__284; +x_386 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__285; lean_inc(x_21); x_387 = l_Lean_Syntax_node1(x_21, x_386, x_385); lean_inc_n(x_387, 2); @@ -7719,7 +7728,7 @@ x_390 = l_Lean_Syntax_node3(x_21, x_28, x_379, x_383, x_389); lean_inc(x_162); lean_inc(x_21); x_391 = l_Lean_Syntax_node2(x_21, x_67, x_162, x_390); -x_392 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__286; +x_392 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__287; lean_inc(x_30); x_393 = lean_array_push(x_392, x_30); lean_inc(x_142); @@ -7731,7 +7740,7 @@ lean_inc(x_30); x_396 = lean_array_push(x_395, x_30); lean_inc(x_364); x_397 = lean_array_push(x_396, x_364); -x_398 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__264; +x_398 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__265; lean_inc(x_397); x_399 = lean_array_push(x_397, x_398); lean_inc(x_36); @@ -7742,18 +7751,18 @@ x_402 = lean_array_push(x_401, x_52); lean_inc(x_79); x_403 = lean_array_push(x_402, x_79); x_404 = lean_array_push(x_403, x_391); -x_405 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__250; +x_405 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__251; lean_inc(x_21); x_406 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_406, 0, x_21); lean_ctor_set(x_406, 1, x_405); lean_ctor_set(x_406, 2, x_404); -x_407 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__292; +x_407 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__293; lean_inc(x_21); x_408 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_408, 0, x_21); lean_ctor_set(x_408, 1, x_407); -x_409 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__291; +x_409 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__292; lean_inc(x_408); lean_inc_n(x_98, 2); lean_inc(x_21); @@ -7764,12 +7773,12 @@ lean_inc(x_44); lean_inc(x_36); lean_inc(x_21); x_411 = l_Lean_Syntax_node5(x_21, x_369, x_36, x_410, x_44, x_66, x_52); -x_412 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__297; +x_412 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__298; lean_inc(x_22); lean_inc(x_23); x_413 = l_Lean_addMacroScope(x_23, x_412, x_22); -x_414 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__294; -x_415 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__299; +x_414 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__295; +x_415 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__300; lean_inc(x_21); x_416 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_416, 0, x_21); @@ -7792,7 +7801,7 @@ x_420 = l_Lean_Syntax_node3(x_21, x_28, x_417, x_383, x_419); lean_inc(x_162); lean_inc(x_21); x_421 = l_Lean_Syntax_node2(x_21, x_67, x_162, x_420); -x_422 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__289; +x_422 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__290; lean_inc(x_397); x_423 = lean_array_push(x_397, x_422); lean_inc(x_36); @@ -7808,12 +7817,12 @@ x_429 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_429, 0, x_21); lean_ctor_set(x_429, 1, x_405); lean_ctor_set(x_429, 2, x_428); -x_430 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__305; +x_430 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__306; lean_inc(x_21); x_431 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_431, 0, x_21); lean_ctor_set(x_431, 1, x_430); -x_432 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__304; +x_432 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__305; lean_inc(x_431); lean_inc_n(x_98, 2); lean_inc(x_21); @@ -7824,12 +7833,12 @@ lean_inc(x_44); lean_inc(x_36); lean_inc(x_21); x_434 = l_Lean_Syntax_node5(x_21, x_369, x_36, x_433, x_44, x_66, x_52); -x_435 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__310; +x_435 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__311; lean_inc(x_22); lean_inc(x_23); x_436 = l_Lean_addMacroScope(x_23, x_435, x_22); -x_437 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__307; -x_438 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__312; +x_437 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__308; +x_438 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__313; lean_inc(x_21); x_439 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_439, 0, x_21); @@ -7852,7 +7861,7 @@ x_443 = l_Lean_Syntax_node3(x_21, x_28, x_440, x_383, x_442); lean_inc(x_162); lean_inc(x_21); x_444 = l_Lean_Syntax_node2(x_21, x_67, x_162, x_443); -x_445 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__302; +x_445 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__303; lean_inc(x_397); x_446 = lean_array_push(x_397, x_445); lean_inc(x_36); @@ -7868,12 +7877,12 @@ x_452 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_452, 0, x_21); lean_ctor_set(x_452, 1, x_405); lean_ctor_set(x_452, 2, x_451); -x_453 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__318; +x_453 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__319; lean_inc(x_21); x_454 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_454, 0, x_21); lean_ctor_set(x_454, 1, x_453); -x_455 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__317; +x_455 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__318; lean_inc(x_454); lean_inc_n(x_98, 2); lean_inc(x_21); @@ -7884,12 +7893,12 @@ lean_inc(x_44); lean_inc(x_36); lean_inc(x_21); x_457 = l_Lean_Syntax_node5(x_21, x_369, x_36, x_456, x_44, x_66, x_52); -x_458 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__323; +x_458 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__324; lean_inc(x_22); lean_inc(x_23); x_459 = l_Lean_addMacroScope(x_23, x_458, x_22); -x_460 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__320; -x_461 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__325; +x_460 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__321; +x_461 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__326; lean_inc(x_21); x_462 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_462, 0, x_21); @@ -7912,7 +7921,7 @@ x_466 = l_Lean_Syntax_node3(x_21, x_28, x_463, x_383, x_465); lean_inc(x_162); lean_inc(x_21); x_467 = l_Lean_Syntax_node2(x_21, x_67, x_162, x_466); -x_468 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__315; +x_468 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__316; lean_inc(x_397); x_469 = lean_array_push(x_397, x_468); lean_inc(x_36); @@ -7928,12 +7937,12 @@ x_475 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_475, 0, x_21); lean_ctor_set(x_475, 1, x_405); lean_ctor_set(x_475, 2, x_474); -x_476 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__331; +x_476 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__332; lean_inc(x_21); x_477 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_477, 0, x_21); lean_ctor_set(x_477, 1, x_476); -x_478 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__330; +x_478 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__331; lean_inc(x_477); lean_inc_n(x_98, 2); lean_inc(x_21); @@ -7944,12 +7953,12 @@ lean_inc(x_44); lean_inc(x_36); lean_inc(x_21); x_480 = l_Lean_Syntax_node5(x_21, x_369, x_36, x_479, x_44, x_66, x_52); -x_481 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__336; +x_481 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__337; lean_inc(x_22); lean_inc(x_23); x_482 = l_Lean_addMacroScope(x_23, x_481, x_22); -x_483 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__333; -x_484 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__338; +x_483 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__334; +x_484 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__339; lean_inc(x_21); x_485 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_485, 0, x_21); @@ -7970,7 +7979,7 @@ lean_inc(x_21); x_489 = l_Lean_Syntax_node3(x_21, x_28, x_486, x_383, x_488); lean_inc(x_21); x_490 = l_Lean_Syntax_node2(x_21, x_67, x_162, x_489); -x_491 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__328; +x_491 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__329; lean_inc(x_397); x_492 = lean_array_push(x_397, x_491); lean_inc(x_36); @@ -7986,7 +7995,7 @@ x_498 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_498, 0, x_21); lean_ctor_set(x_498, 1, x_405); lean_ctor_set(x_498, 2, x_497); -x_499 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__341; +x_499 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__342; lean_inc(x_21); x_500 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_500, 0, x_21); @@ -7998,23 +8007,23 @@ lean_inc(x_98); lean_inc(x_36); lean_inc(x_21); x_501 = l_Lean_Syntax_node5(x_21, x_369, x_36, x_98, x_44, x_66, x_52); -x_502 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__347; +x_502 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__348; lean_inc(x_21); x_503 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_503, 0, x_21); lean_ctor_set(x_503, 1, x_502); -x_504 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__346; +x_504 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__347; lean_inc(x_98); lean_inc(x_503); lean_inc(x_501); lean_inc(x_21); x_505 = l_Lean_Syntax_node3(x_21, x_504, x_501, x_503, x_98); -x_506 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__352; +x_506 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__353; lean_inc(x_22); lean_inc(x_23); x_507 = l_Lean_addMacroScope(x_23, x_506, x_22); -x_508 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__349; -x_509 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__354; +x_508 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__350; +x_509 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__355; lean_inc(x_21); x_510 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_510, 0, x_21); @@ -8024,7 +8033,7 @@ lean_ctor_set(x_510, 3, x_509); lean_inc_n(x_372, 2); lean_inc(x_21); x_511 = l_Lean_Syntax_node3(x_21, x_378, x_372, x_372, x_510); -x_512 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__355; +x_512 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__356; lean_inc(x_21); x_513 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_513, 0, x_21); @@ -8050,7 +8059,7 @@ x_520 = lean_array_push(x_394, x_500); lean_inc(x_30); x_521 = lean_array_push(x_520, x_30); x_522 = lean_array_push(x_521, x_364); -x_523 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__344; +x_523 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__345; lean_inc(x_522); x_524 = lean_array_push(x_522, x_523); lean_inc(x_36); @@ -8061,29 +8070,29 @@ x_527 = lean_array_push(x_526, x_52); lean_inc(x_79); x_528 = lean_array_push(x_527, x_79); x_529 = lean_array_push(x_528, x_519); -x_530 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__340; +x_530 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__341; lean_inc(x_21); x_531 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_531, 0, x_21); lean_ctor_set(x_531, 1, x_530); lean_ctor_set(x_531, 2, x_529); -x_532 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__361; +x_532 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__362; lean_inc(x_21); x_533 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_533, 0, x_21); lean_ctor_set(x_533, 1, x_532); -x_534 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__360; +x_534 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__361; lean_inc(x_98); lean_inc(x_533); lean_inc(x_501); lean_inc(x_21); x_535 = l_Lean_Syntax_node3(x_21, x_534, x_501, x_533, x_98); -x_536 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__366; +x_536 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__367; lean_inc(x_22); lean_inc(x_23); x_537 = l_Lean_addMacroScope(x_23, x_536, x_22); -x_538 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__363; -x_539 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__368; +x_538 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__364; +x_539 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__369; lean_inc(x_21); x_540 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_540, 0, x_21); @@ -8106,7 +8115,7 @@ x_544 = l_Lean_Syntax_node3(x_21, x_28, x_541, x_514, x_543); lean_inc(x_305); lean_inc(x_21); x_545 = l_Lean_Syntax_node2(x_21, x_67, x_305, x_544); -x_546 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__358; +x_546 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__359; lean_inc(x_522); x_547 = lean_array_push(x_522, x_546); lean_inc(x_36); @@ -8122,23 +8131,23 @@ x_553 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_553, 0, x_21); lean_ctor_set(x_553, 1, x_530); lean_ctor_set(x_553, 2, x_552); -x_554 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__374; +x_554 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__375; lean_inc(x_21); x_555 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_555, 0, x_21); lean_ctor_set(x_555, 1, x_554); -x_556 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__373; +x_556 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__374; lean_inc(x_98); lean_inc(x_555); lean_inc(x_501); lean_inc(x_21); x_557 = l_Lean_Syntax_node3(x_21, x_556, x_501, x_555, x_98); -x_558 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__379; +x_558 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__380; lean_inc(x_22); lean_inc(x_23); x_559 = l_Lean_addMacroScope(x_23, x_558, x_22); -x_560 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__376; -x_561 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__381; +x_560 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__377; +x_561 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__382; lean_inc(x_21); x_562 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_562, 0, x_21); @@ -8161,7 +8170,7 @@ x_566 = l_Lean_Syntax_node3(x_21, x_28, x_563, x_514, x_565); lean_inc(x_305); lean_inc(x_21); x_567 = l_Lean_Syntax_node2(x_21, x_67, x_305, x_566); -x_568 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__371; +x_568 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__372; lean_inc(x_522); x_569 = lean_array_push(x_522, x_568); lean_inc(x_36); @@ -8177,23 +8186,23 @@ x_575 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_575, 0, x_21); lean_ctor_set(x_575, 1, x_530); lean_ctor_set(x_575, 2, x_574); -x_576 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__387; +x_576 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__388; lean_inc(x_21); x_577 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_577, 0, x_21); lean_ctor_set(x_577, 1, x_576); -x_578 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__386; +x_578 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__387; lean_inc(x_98); lean_inc(x_577); lean_inc(x_501); lean_inc(x_21); x_579 = l_Lean_Syntax_node3(x_21, x_578, x_501, x_577, x_98); -x_580 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__392; +x_580 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__393; lean_inc(x_22); lean_inc(x_23); x_581 = l_Lean_addMacroScope(x_23, x_580, x_22); -x_582 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__389; -x_583 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__394; +x_582 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__390; +x_583 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__395; lean_inc(x_21); x_584 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_584, 0, x_21); @@ -8216,7 +8225,7 @@ x_588 = l_Lean_Syntax_node3(x_21, x_28, x_585, x_514, x_587); lean_inc(x_305); lean_inc(x_21); x_589 = l_Lean_Syntax_node2(x_21, x_67, x_305, x_588); -x_590 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__384; +x_590 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__385; lean_inc(x_522); x_591 = lean_array_push(x_522, x_590); lean_inc(x_36); @@ -8232,34 +8241,34 @@ x_597 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_597, 0, x_21); lean_ctor_set(x_597, 1, x_530); lean_ctor_set(x_597, 2, x_596); -x_598 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__397; +x_598 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__398; lean_inc(x_22); lean_inc(x_23); x_599 = l_Lean_addMacroScope(x_23, x_598, x_22); -x_600 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__396; +x_600 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__397; lean_inc(x_21); x_601 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_601, 0, x_21); lean_ctor_set(x_601, 1, x_600); lean_ctor_set(x_601, 2, x_599); lean_ctor_set(x_601, 3, x_39); -x_602 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__400; +x_602 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__401; lean_inc(x_21); x_603 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_603, 0, x_21); lean_ctor_set(x_603, 1, x_602); -x_604 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__399; +x_604 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__400; lean_inc(x_98); lean_inc(x_603); lean_inc(x_501); lean_inc(x_21); x_605 = l_Lean_Syntax_node3(x_21, x_604, x_501, x_603, x_98); -x_606 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__403; +x_606 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__404; lean_inc(x_22); lean_inc(x_23); x_607 = l_Lean_addMacroScope(x_23, x_606, x_22); -x_608 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__402; -x_609 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__407; +x_608 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__403; +x_609 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__408; lean_inc(x_21); x_610 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_610, 0, x_21); @@ -8269,7 +8278,7 @@ lean_ctor_set(x_610, 3, x_609); lean_inc_n(x_372, 2); lean_inc(x_21); x_611 = l_Lean_Syntax_node3(x_21, x_378, x_372, x_372, x_610); -x_612 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__408; +x_612 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__409; lean_inc(x_21); x_613 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_613, 0, x_21); @@ -8304,34 +8313,34 @@ x_625 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_625, 0, x_21); lean_ctor_set(x_625, 1, x_530); lean_ctor_set(x_625, 2, x_624); -x_626 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__411; +x_626 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__412; lean_inc(x_22); lean_inc(x_23); x_627 = l_Lean_addMacroScope(x_23, x_626, x_22); -x_628 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__410; +x_628 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__411; lean_inc(x_21); x_629 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_629, 0, x_21); lean_ctor_set(x_629, 1, x_628); lean_ctor_set(x_629, 2, x_627); lean_ctor_set(x_629, 3, x_39); -x_630 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__414; +x_630 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__415; lean_inc(x_21); x_631 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_631, 0, x_21); lean_ctor_set(x_631, 1, x_630); -x_632 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__413; +x_632 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__414; lean_inc(x_98); lean_inc(x_631); lean_inc(x_501); lean_inc(x_21); x_633 = l_Lean_Syntax_node3(x_21, x_632, x_501, x_631, x_98); -x_634 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__417; +x_634 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__418; lean_inc(x_22); lean_inc(x_23); x_635 = l_Lean_addMacroScope(x_23, x_634, x_22); -x_636 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__416; -x_637 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__421; +x_636 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__417; +x_637 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__422; lean_inc(x_21); x_638 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_638, 0, x_21); @@ -8367,34 +8376,34 @@ x_650 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_650, 0, x_21); lean_ctor_set(x_650, 1, x_530); lean_ctor_set(x_650, 2, x_649); -x_651 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__424; +x_651 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__425; lean_inc(x_22); lean_inc(x_23); x_652 = l_Lean_addMacroScope(x_23, x_651, x_22); -x_653 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__423; +x_653 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__424; lean_inc(x_21); x_654 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_654, 0, x_21); lean_ctor_set(x_654, 1, x_653); lean_ctor_set(x_654, 2, x_652); lean_ctor_set(x_654, 3, x_39); -x_655 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__427; +x_655 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__428; lean_inc(x_21); x_656 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_656, 0, x_21); lean_ctor_set(x_656, 1, x_655); -x_657 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__426; +x_657 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__427; lean_inc(x_98); lean_inc(x_656); lean_inc(x_501); lean_inc(x_21); x_658 = l_Lean_Syntax_node3(x_21, x_657, x_501, x_656, x_98); -x_659 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__432; +x_659 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__433; lean_inc(x_22); lean_inc(x_23); x_660 = l_Lean_addMacroScope(x_23, x_659, x_22); -x_661 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__429; -x_662 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__434; +x_661 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__430; +x_662 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__435; lean_inc(x_21); x_663 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_663, 0, x_21); @@ -8432,33 +8441,33 @@ x_675 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_675, 0, x_21); lean_ctor_set(x_675, 1, x_405); lean_ctor_set(x_675, 2, x_674); -x_676 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__437; +x_676 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__438; lean_inc(x_22); lean_inc(x_23); x_677 = l_Lean_addMacroScope(x_23, x_676, x_22); -x_678 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__436; +x_678 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__437; lean_inc(x_21); x_679 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_679, 0, x_21); lean_ctor_set(x_679, 1, x_678); lean_ctor_set(x_679, 2, x_677); lean_ctor_set(x_679, 3, x_39); -x_680 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__440; +x_680 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__441; lean_inc(x_21); x_681 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_681, 0, x_21); lean_ctor_set(x_681, 1, x_680); -x_682 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__439; +x_682 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__440; lean_inc(x_98); lean_inc(x_681); lean_inc(x_21); x_683 = l_Lean_Syntax_node3(x_21, x_682, x_501, x_681, x_98); -x_684 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__443; +x_684 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__444; lean_inc(x_22); lean_inc(x_23); x_685 = l_Lean_addMacroScope(x_23, x_684, x_22); -x_686 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__442; -x_687 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__447; +x_686 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__443; +x_687 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__448; lean_inc(x_21); x_688 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_688, 0, x_21); @@ -8500,12 +8509,12 @@ x_701 = l_Lean_Syntax_node2(x_21, x_28, x_98, x_98); lean_inc(x_15); lean_inc(x_21); x_702 = l_Lean_Syntax_node2(x_21, x_67, x_15, x_701); -x_703 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__451; +x_703 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__452; lean_inc(x_21); x_704 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_704, 0, x_21); lean_ctor_set(x_704, 1, x_703); -x_705 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__455; +x_705 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__456; lean_inc(x_21); x_706 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_706, 0, x_21); @@ -8514,18 +8523,18 @@ x_707 = l_Lean_Syntax_getId(x_15); lean_dec(x_15); lean_inc(x_707); x_708 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_39, x_707); -x_709 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__456; +x_709 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__457; lean_inc(x_21); x_710 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_710, 0, x_21); lean_ctor_set(x_710, 1, x_709); lean_inc(x_21); x_711 = l_Lean_Syntax_node1(x_21, x_382, x_710); -x_712 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__459; +x_712 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__460; lean_inc(x_22); lean_inc(x_23); x_713 = l_Lean_addMacroScope(x_23, x_712, x_22); -x_714 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__458; +x_714 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__459; lean_inc(x_21); x_715 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_715, 0, x_21); @@ -8539,12 +8548,12 @@ lean_inc(x_716); lean_inc(x_88); lean_inc(x_21); x_717 = l_Lean_Syntax_node2(x_21, x_67, x_88, x_716); -x_718 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__463; +x_718 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__464; lean_inc(x_22); lean_inc(x_23); x_719 = l_Lean_addMacroScope(x_23, x_718, x_22); -x_720 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__461; -x_721 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__467; +x_720 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__462; +x_721 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__468; lean_inc(x_21); x_722 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_722, 0, x_21); @@ -8572,15 +8581,15 @@ lean_inc(x_716); lean_inc(x_131); lean_inc(x_21); x_728 = l_Lean_Syntax_node2(x_21, x_67, x_131, x_716); -x_729 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__473; +x_729 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__474; lean_inc(x_79); lean_inc_n(x_30, 2); lean_inc(x_21); x_730 = l_Lean_Syntax_node5(x_21, x_729, x_715, x_30, x_30, x_79, x_728); -x_731 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__471; +x_731 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__472; lean_inc(x_21); x_732 = l_Lean_Syntax_node1(x_21, x_731, x_730); -x_733 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__469; +x_733 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__470; lean_inc(x_30); lean_inc(x_83); lean_inc(x_21); @@ -8603,7 +8612,7 @@ x_739 = l_Lean_Syntax_node2(x_21, x_123, x_116, x_738); lean_inc(x_30); lean_inc(x_21); x_740 = l_Lean_Syntax_node2(x_21, x_125, x_739, x_30); -x_741 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__450; +x_741 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__451; lean_inc(x_397); x_742 = lean_array_push(x_397, x_741); lean_inc(x_36); @@ -8618,7 +8627,7 @@ lean_inc(x_21); x_747 = l_Lean_Syntax_node2(x_21, x_67, x_131, x_99); lean_inc(x_12); x_748 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_39, x_12); -x_749 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__477; +x_749 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__478; lean_inc(x_21); x_750 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_750, 0, x_21); @@ -8640,7 +8649,7 @@ x_755 = l_Lean_Syntax_node3(x_21, x_250, x_83, x_30, x_754); lean_inc(x_30); lean_inc(x_21); x_756 = l_Lean_Syntax_node2(x_21, x_125, x_755, x_30); -x_757 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__476; +x_757 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__477; lean_inc(x_397); x_758 = lean_array_push(x_397, x_757); lean_inc(x_36); @@ -8657,11 +8666,11 @@ x_763 = l_Lean_Syntax_node2(x_21, x_67, x_18, x_99); x_764 = l_Lean_Syntax_getId(x_18); lean_inc(x_764); x_765 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_39, x_764); -x_766 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__483; +x_766 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__484; lean_inc(x_22); lean_inc(x_23); x_767 = l_Lean_addMacroScope(x_23, x_766, x_22); -x_768 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__482; +x_768 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__483; lean_inc(x_21); x_769 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_769, 0, x_21); @@ -8711,7 +8720,7 @@ x_783 = l_Lean_Syntax_node2(x_21, x_123, x_116, x_782); lean_inc(x_30); lean_inc(x_21); x_784 = l_Lean_Syntax_node2(x_21, x_125, x_783, x_30); -x_785 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__480; +x_785 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__481; x_786 = lean_array_push(x_397, x_785); lean_inc(x_36); x_787 = lean_array_push(x_786, x_36); @@ -8720,17 +8729,17 @@ lean_inc(x_52); x_789 = lean_array_push(x_788, x_52); lean_inc(x_79); x_790 = lean_array_push(x_789, x_79); -x_791 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__486; +x_791 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__487; lean_inc(x_21); x_792 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_792, 0, x_21); lean_ctor_set(x_792, 1, x_791); -x_793 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__487; +x_793 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__488; lean_inc(x_21); x_794 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_794, 0, x_21); lean_ctor_set(x_794, 1, x_793); -x_795 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__485; +x_795 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__486; lean_inc(x_21); x_796 = l_Lean_Syntax_node2(x_21, x_795, x_792, x_794); lean_inc(x_21); @@ -8739,21 +8748,21 @@ lean_inc(x_21); x_798 = l_Lean_Syntax_node1(x_21, x_28, x_362); lean_inc(x_21); x_799 = l_Lean_Syntax_node3(x_21, x_28, x_353, x_798, x_154); -x_800 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__490; +x_800 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__491; lean_inc(x_22); lean_inc(x_23); x_801 = l_Lean_addMacroScope(x_23, x_800, x_22); -x_802 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__489; +x_802 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__490; lean_inc(x_21); x_803 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_803, 0, x_21); lean_ctor_set(x_803, 1, x_802); lean_ctor_set(x_803, 2, x_801); lean_ctor_set(x_803, 3, x_39); -x_804 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__494; +x_804 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__495; x_805 = l_Lean_addMacroScope(x_23, x_804, x_22); -x_806 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__492; -x_807 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__496; +x_806 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__493; +x_807 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__497; lean_inc(x_21); x_808 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_808, 0, x_21); @@ -8805,13 +8814,13 @@ x_822 = l_Lean_Syntax_node1(x_21, x_128, x_821); lean_inc(x_81); lean_inc(x_21); x_823 = l_Lean_Syntax_node2(x_21, x_294, x_81, x_822); -x_824 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__454; +x_824 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__455; lean_inc(x_706); lean_inc(x_30); lean_inc(x_42); lean_inc(x_21); x_825 = l_Lean_Syntax_node4(x_21, x_824, x_42, x_30, x_706, x_823); -x_826 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__452; +x_826 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__453; lean_inc(x_704); lean_inc(x_21); x_827 = l_Lean_Syntax_node2(x_21, x_826, x_704, x_825); @@ -8833,15 +8842,15 @@ x_839 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_839, 0, x_21); lean_ctor_set(x_839, 1, x_405); lean_ctor_set(x_839, 2, x_838); -x_840 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__497; +x_840 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__498; lean_inc(x_21); x_841 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_841, 0, x_21); lean_ctor_set(x_841, 1, x_840); -x_842 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__498; +x_842 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__499; lean_inc(x_21); x_843 = l_Lean_Syntax_node2(x_21, x_842, x_841, x_66); -x_844 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__499; +x_844 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__500; x_845 = lean_array_push(x_844, x_27); if (lean_obj_tag(x_112) == 0) { @@ -8862,9 +8871,9 @@ x_971 = lean_string_append(x_371, x_970); lean_dec(x_970); x_972 = lean_box(2); x_973 = l_Lean_Syntax_mkNameLit(x_971, x_972); -x_974 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__502; +x_974 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__503; x_975 = lean_array_push(x_974, x_973); -x_976 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__501; +x_976 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__502; x_977 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_977, 0, x_972); lean_ctor_set(x_977, 1, x_976); @@ -8899,7 +8908,7 @@ x_855 = l_Lean_Syntax_node2(x_21, x_294, x_81, x_854); lean_inc(x_30); lean_inc(x_21); x_856 = l_Lean_Syntax_node4(x_21, x_296, x_79, x_855, x_138, x_30); -x_857 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__30; +x_857 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__31; lean_inc(x_30); lean_inc(x_21); x_858 = l_Lean_Syntax_node5(x_21, x_298, x_34, x_857, x_77, x_856, x_30); @@ -8941,9 +8950,9 @@ x_960 = lean_string_append(x_371, x_959); lean_dec(x_959); x_961 = lean_box(2); x_962 = l_Lean_Syntax_mkNameLit(x_960, x_961); -x_963 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__502; +x_963 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__503; x_964 = lean_array_push(x_963, x_962); -x_965 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__501; +x_965 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__502; x_966 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_966, 0, x_961); lean_ctor_set(x_966, 1, x_965); @@ -9010,9 +9019,9 @@ x_949 = lean_string_append(x_371, x_948); lean_dec(x_948); x_950 = lean_box(2); x_951 = l_Lean_Syntax_mkNameLit(x_949, x_950); -x_952 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__502; +x_952 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__503; x_953 = lean_array_push(x_952, x_951); -x_954 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__501; +x_954 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__502; x_955 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_955, 0, x_950); lean_ctor_set(x_955, 1, x_954); @@ -9113,9 +9122,9 @@ x_922 = lean_string_append(x_371, x_921); lean_dec(x_921); x_923 = lean_box(2); x_924 = l_Lean_Syntax_mkNameLit(x_922, x_923); -x_925 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__502; +x_925 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__503; x_926 = lean_array_push(x_925, x_924); -x_927 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__501; +x_927 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__502; x_928 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_928, 0, x_923); lean_ctor_set(x_928, 1, x_927); @@ -9367,7 +9376,7 @@ static lean_object* _init_l_UInt8_reduceBin____x40_Lean_Meta_Tactic_Simp_Builtin _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__494; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__495; x_2 = l_UInt8_reduceBin____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_7772____lambda__1___closed__3; x_3 = l_Lean_Expr_const___override(x_1, x_2); return x_3; @@ -9917,7 +9926,7 @@ static lean_object* _init_l_UInt8_reduceBoolPred____x40_Lean_Meta_Tactic_Simp_Bu _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__223; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__224; x_2 = l_UInt8_reduceBoolPred____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_7772____lambda__1___closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; @@ -9955,7 +9964,7 @@ static lean_object* _init_l_UInt8_reduceBoolPred____x40_Lean_Meta_Tactic_Simp_Bu _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__223; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__224; x_2 = l_UInt8_reduceBoolPred____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_7772____lambda__1___closed__5; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; @@ -10455,7 +10464,7 @@ LEAN_EXPORT lean_object* l_UInt8_reduceAdd(lean_object* x_1, lean_object* x_2, l _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__277; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__278; x_11 = lean_unsigned_to_nat(6u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -10504,7 +10513,7 @@ static lean_object* _init_l___regBuiltin_UInt8_reduceAdd_declare____x40_Lean_Met { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt8_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__262; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__263; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -10513,7 +10522,7 @@ static lean_object* _init_l___regBuiltin_UInt8_reduceAdd_declare____x40_Lean_Met _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__277; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__278; x_2 = lean_unsigned_to_nat(6u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10893,7 +10902,7 @@ LEAN_EXPORT lean_object* l_UInt8_reduceMul(lean_object* x_1, lean_object* x_2, l _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__297; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__298; x_11 = lean_unsigned_to_nat(6u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -10942,7 +10951,7 @@ static lean_object* _init_l___regBuiltin_UInt8_reduceMul_declare____x40_Lean_Met { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt8_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__287; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__288; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -10951,7 +10960,7 @@ static lean_object* _init_l___regBuiltin_UInt8_reduceMul_declare____x40_Lean_Met _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__297; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__298; x_2 = lean_unsigned_to_nat(6u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11294,7 +11303,7 @@ LEAN_EXPORT lean_object* l_UInt8_reduceSub(lean_object* x_1, lean_object* x_2, l _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__310; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__311; x_11 = lean_unsigned_to_nat(6u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -11343,7 +11352,7 @@ static lean_object* _init_l___regBuiltin_UInt8_reduceSub_declare____x40_Lean_Met { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt8_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__300; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__301; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -11352,7 +11361,7 @@ static lean_object* _init_l___regBuiltin_UInt8_reduceSub_declare____x40_Lean_Met _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__310; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__311; x_2 = lean_unsigned_to_nat(6u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11695,7 +11704,7 @@ LEAN_EXPORT lean_object* l_UInt8_reduceDiv(lean_object* x_1, lean_object* x_2, l _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__323; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__324; x_11 = lean_unsigned_to_nat(6u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -11744,7 +11753,7 @@ static lean_object* _init_l___regBuiltin_UInt8_reduceDiv_declare____x40_Lean_Met { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt8_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__313; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__314; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -11753,7 +11762,7 @@ static lean_object* _init_l___regBuiltin_UInt8_reduceDiv_declare____x40_Lean_Met _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__323; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__324; x_2 = lean_unsigned_to_nat(6u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -12096,7 +12105,7 @@ LEAN_EXPORT lean_object* l_UInt8_reduceMod(lean_object* x_1, lean_object* x_2, l _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__336; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__337; x_11 = lean_unsigned_to_nat(6u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -12145,7 +12154,7 @@ static lean_object* _init_l___regBuiltin_UInt8_reduceMod_declare____x40_Lean_Met { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt8_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__326; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__327; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -12154,7 +12163,7 @@ static lean_object* _init_l___regBuiltin_UInt8_reduceMod_declare____x40_Lean_Met _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__336; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__337; x_2 = lean_unsigned_to_nat(6u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -12467,7 +12476,7 @@ LEAN_EXPORT lean_object* l_UInt8_reduceLT(lean_object* x_1, lean_object* x_2, le _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__352; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__353; x_11 = lean_unsigned_to_nat(4u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -12516,7 +12525,7 @@ static lean_object* _init_l___regBuiltin_UInt8_reduceLT_declare____x40_Lean_Meta { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt8_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__342; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__343; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -12525,7 +12534,7 @@ static lean_object* _init_l___regBuiltin_UInt8_reduceLT_declare____x40_Lean_Meta _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__352; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__353; x_2 = lean_unsigned_to_nat(4u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -12827,7 +12836,7 @@ LEAN_EXPORT lean_object* l_UInt8_reduceLE(lean_object* x_1, lean_object* x_2, le _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__366; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__367; x_11 = lean_unsigned_to_nat(4u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -12876,7 +12885,7 @@ static lean_object* _init_l___regBuiltin_UInt8_reduceLE_declare____x40_Lean_Meta { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt8_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__356; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__357; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -12885,7 +12894,7 @@ static lean_object* _init_l___regBuiltin_UInt8_reduceLE_declare____x40_Lean_Meta _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__366; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__367; x_2 = lean_unsigned_to_nat(4u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13178,7 +13187,7 @@ LEAN_EXPORT lean_object* l_UInt8_reduceGT(lean_object* x_1, lean_object* x_2, le _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__379; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__380; x_11 = lean_unsigned_to_nat(4u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -13227,7 +13236,7 @@ static lean_object* _init_l___regBuiltin_UInt8_reduceGT_declare____x40_Lean_Meta { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt8_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__369; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__370; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -13467,7 +13476,7 @@ LEAN_EXPORT lean_object* l_UInt8_reduceGE(lean_object* x_1, lean_object* x_2, le _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__392; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__393; x_11 = lean_unsigned_to_nat(4u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -13516,7 +13525,7 @@ static lean_object* _init_l___regBuiltin_UInt8_reduceGE_declare____x40_Lean_Meta { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt8_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__382; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__383; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -13756,7 +13765,7 @@ LEAN_EXPORT lean_object* l_UInt8_reduceEq____x40_Lean_Meta_Tactic_Simp_BuiltinSi _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__403; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__404; x_11 = lean_unsigned_to_nat(3u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -13805,7 +13814,7 @@ static lean_object* _init_l___regBuiltin_UInt8_reduceEq_declare____x40_Lean_Meta { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt8_fromExpr___closed__2; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__395; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__396; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -13843,7 +13852,7 @@ static lean_object* _init_l___regBuiltin_UInt8_reduceEq_declare____x40_Lean_Meta { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_UInt8_reduceEq_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_7772__8669____closed__4; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__57; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -13871,7 +13880,7 @@ static lean_object* _init_l___regBuiltin_UInt8_reduceEq_declare____x40_Lean_Meta { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_UInt8_reduceEq_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_7772__8669____closed__7; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__59; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -13944,7 +13953,7 @@ static lean_object* _init_l___regBuiltin_UInt8_reduceEq_declare____x40_Lean_Meta _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__403; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__404; x_2 = lean_unsigned_to_nat(3u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -14237,7 +14246,7 @@ LEAN_EXPORT lean_object* l_UInt8_reduceNe____x40_Lean_Meta_Tactic_Simp_BuiltinSi _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__417; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__418; x_11 = lean_unsigned_to_nat(3u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -14286,7 +14295,7 @@ static lean_object* _init_l___regBuiltin_UInt8_reduceNe_declare____x40_Lean_Meta { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt8_fromExpr___closed__2; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__409; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__410; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -14316,7 +14325,7 @@ static lean_object* _init_l___regBuiltin_UInt8_reduceNe_declare____x40_Lean_Meta { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_UInt8_reduceNe_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_7772__8709____closed__3; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__57; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -14336,7 +14345,7 @@ static lean_object* _init_l___regBuiltin_UInt8_reduceNe_declare____x40_Lean_Meta { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_UInt8_reduceNe_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_7772__8709____closed__5; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__59; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -14732,7 +14741,7 @@ LEAN_EXPORT lean_object* l_UInt8_reduceBEq____x40_Lean_Meta_Tactic_Simp_BuiltinS _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__432; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__433; x_11 = lean_unsigned_to_nat(4u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -14781,7 +14790,7 @@ static lean_object* _init_l___regBuiltin_UInt8_reduceBEq_declare____x40_Lean_Met { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt8_fromExpr___closed__2; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__422; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__423; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -14811,7 +14820,7 @@ static lean_object* _init_l___regBuiltin_UInt8_reduceBEq_declare____x40_Lean_Met { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_UInt8_reduceBEq_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_7772__8750____closed__3; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__57; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -14831,7 +14840,7 @@ static lean_object* _init_l___regBuiltin_UInt8_reduceBEq_declare____x40_Lean_Met { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_UInt8_reduceBEq_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_7772__8750____closed__5; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__59; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -14880,7 +14889,7 @@ static lean_object* _init_l___regBuiltin_UInt8_reduceBEq_declare____x40_Lean_Met _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__432; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__433; x_2 = lean_unsigned_to_nat(4u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -15209,7 +15218,7 @@ LEAN_EXPORT lean_object* l_UInt8_reduceBNe____x40_Lean_Meta_Tactic_Simp_BuiltinS _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__443; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__444; x_11 = lean_unsigned_to_nat(4u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -15258,7 +15267,7 @@ static lean_object* _init_l___regBuiltin_UInt8_reduceBNe_declare____x40_Lean_Met { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt8_fromExpr___closed__2; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__435; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__436; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -15288,7 +15297,7 @@ static lean_object* _init_l___regBuiltin_UInt8_reduceBNe_declare____x40_Lean_Met { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_UInt8_reduceBNe_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_7772__8790____closed__3; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__57; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -15308,7 +15317,7 @@ static lean_object* _init_l___regBuiltin_UInt8_reduceBNe_declare____x40_Lean_Met { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_UInt8_reduceBNe_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_7772__8790____closed__5; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__59; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -15357,7 +15366,7 @@ static lean_object* _init_l___regBuiltin_UInt8_reduceBNe_declare____x40_Lean_Met _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__443; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__444; x_2 = lean_unsigned_to_nat(4u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -15649,7 +15658,7 @@ static lean_object* _init_l___regBuiltin_UInt8_reduceOfNatCore_declare____x40_Le { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt8_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__448; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__449; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -15946,7 +15955,7 @@ static lean_object* _init_l___regBuiltin_UInt8_reduceOfNat_declare____x40_Lean_M { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt8_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__474; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__475; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -15967,7 +15976,7 @@ static lean_object* _init_l___regBuiltin_UInt8_reduceOfNat_declare____x40_Lean_M _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__27; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__28; x_2 = l___regBuiltin_UInt8_reduceOfNat_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_9017____closed__2; x_3 = lean_array_push(x_1, x_2); return x_3; @@ -16212,7 +16221,7 @@ static lean_object* _init_l___regBuiltin_UInt8_reduceToNat_declare____x40_Lean_M { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt8_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__478; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__479; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -16233,7 +16242,7 @@ static lean_object* _init_l___regBuiltin_UInt8_reduceToNat_declare____x40_Lean_M _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__27; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__28; x_2 = l___regBuiltin_UInt8_reduceToNat_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_9132____closed__2; x_3 = lean_array_push(x_1, x_2); return x_3; @@ -16318,7 +16327,7 @@ LEAN_EXPORT lean_object* l_UInt8_isValue____x40_Lean_Meta_Tactic_Simp_BuiltinSim _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__494; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__495; x_11 = lean_unsigned_to_nat(3u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -16375,7 +16384,7 @@ static lean_object* _init_l___regBuiltin_UInt8_isValue_declare____x40_Lean_Meta_ { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt8_fromExpr___closed__2; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__488; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__489; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -16405,7 +16414,7 @@ static lean_object* _init_l___regBuiltin_UInt8_isValue_declare____x40_Lean_Meta_ { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_UInt8_isValue_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_7772__9205____closed__3; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__57; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -16425,7 +16434,7 @@ static lean_object* _init_l___regBuiltin_UInt8_isValue_declare____x40_Lean_Meta_ { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_UInt8_isValue_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_7772__9205____closed__5; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__59; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -16474,7 +16483,7 @@ static lean_object* _init_l___regBuiltin_UInt8_isValue_declare____x40_Lean_Meta_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__494; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__495; x_2 = lean_unsigned_to_nat(3u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -17719,7 +17728,7 @@ LEAN_EXPORT lean_object* l_UInt16_reduceAdd(lean_object* x_1, lean_object* x_2, _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__277; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__278; x_11 = lean_unsigned_to_nat(6u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -17768,7 +17777,7 @@ static lean_object* _init_l___regBuiltin_UInt16_reduceAdd_declare____x40_Lean_Me { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt16_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__262; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__263; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -18110,7 +18119,7 @@ LEAN_EXPORT lean_object* l_UInt16_reduceMul(lean_object* x_1, lean_object* x_2, _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__297; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__298; x_11 = lean_unsigned_to_nat(6u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -18159,7 +18168,7 @@ static lean_object* _init_l___regBuiltin_UInt16_reduceMul_declare____x40_Lean_Me { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt16_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__287; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__288; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -18489,7 +18498,7 @@ LEAN_EXPORT lean_object* l_UInt16_reduceSub(lean_object* x_1, lean_object* x_2, _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__310; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__311; x_11 = lean_unsigned_to_nat(6u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -18538,7 +18547,7 @@ static lean_object* _init_l___regBuiltin_UInt16_reduceSub_declare____x40_Lean_Me { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt16_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__300; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__301; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -18868,7 +18877,7 @@ LEAN_EXPORT lean_object* l_UInt16_reduceDiv(lean_object* x_1, lean_object* x_2, _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__323; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__324; x_11 = lean_unsigned_to_nat(6u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -18917,7 +18926,7 @@ static lean_object* _init_l___regBuiltin_UInt16_reduceDiv_declare____x40_Lean_Me { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt16_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__313; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__314; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -19247,7 +19256,7 @@ LEAN_EXPORT lean_object* l_UInt16_reduceMod(lean_object* x_1, lean_object* x_2, _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__336; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__337; x_11 = lean_unsigned_to_nat(6u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -19296,7 +19305,7 @@ static lean_object* _init_l___regBuiltin_UInt16_reduceMod_declare____x40_Lean_Me { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt16_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__326; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__327; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -19596,7 +19605,7 @@ LEAN_EXPORT lean_object* l_UInt16_reduceLT(lean_object* x_1, lean_object* x_2, l _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__352; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__353; x_11 = lean_unsigned_to_nat(4u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -19645,7 +19654,7 @@ static lean_object* _init_l___regBuiltin_UInt16_reduceLT_declare____x40_Lean_Met { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt16_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__342; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__343; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -19925,7 +19934,7 @@ LEAN_EXPORT lean_object* l_UInt16_reduceLE(lean_object* x_1, lean_object* x_2, l _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__366; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__367; x_11 = lean_unsigned_to_nat(4u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -19974,7 +19983,7 @@ static lean_object* _init_l___regBuiltin_UInt16_reduceLE_declare____x40_Lean_Met { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt16_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__356; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__357; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -20254,7 +20263,7 @@ LEAN_EXPORT lean_object* l_UInt16_reduceGT(lean_object* x_1, lean_object* x_2, l _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__379; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__380; x_11 = lean_unsigned_to_nat(4u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -20303,7 +20312,7 @@ static lean_object* _init_l___regBuiltin_UInt16_reduceGT_declare____x40_Lean_Met { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt16_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__369; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__370; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -20543,7 +20552,7 @@ LEAN_EXPORT lean_object* l_UInt16_reduceGE(lean_object* x_1, lean_object* x_2, l _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__392; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__393; x_11 = lean_unsigned_to_nat(4u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -20592,7 +20601,7 @@ static lean_object* _init_l___regBuiltin_UInt16_reduceGE_declare____x40_Lean_Met { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt16_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__382; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__383; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -20832,7 +20841,7 @@ LEAN_EXPORT lean_object* l_UInt16_reduceEq____x40_Lean_Meta_Tactic_Simp_BuiltinS _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__403; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__404; x_11 = lean_unsigned_to_nat(3u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -20881,7 +20890,7 @@ static lean_object* _init_l___regBuiltin_UInt16_reduceEq_declare____x40_Lean_Met { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt16_fromExpr___closed__2; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__395; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__396; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -20911,7 +20920,7 @@ static lean_object* _init_l___regBuiltin_UInt16_reduceEq_declare____x40_Lean_Met { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_UInt16_reduceEq_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_9210__10107____closed__3; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__57; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -20931,7 +20940,7 @@ static lean_object* _init_l___regBuiltin_UInt16_reduceEq_declare____x40_Lean_Met { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_UInt16_reduceEq_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_9210__10107____closed__5; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__59; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -21242,7 +21251,7 @@ LEAN_EXPORT lean_object* l_UInt16_reduceNe____x40_Lean_Meta_Tactic_Simp_BuiltinS _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__417; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__418; x_11 = lean_unsigned_to_nat(3u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -21291,7 +21300,7 @@ static lean_object* _init_l___regBuiltin_UInt16_reduceNe_declare____x40_Lean_Met { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt16_fromExpr___closed__2; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__409; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__410; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -21321,7 +21330,7 @@ static lean_object* _init_l___regBuiltin_UInt16_reduceNe_declare____x40_Lean_Met { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_UInt16_reduceNe_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_9210__10147____closed__3; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__57; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -21341,7 +21350,7 @@ static lean_object* _init_l___regBuiltin_UInt16_reduceNe_declare____x40_Lean_Met { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_UInt16_reduceNe_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_9210__10147____closed__5; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__59; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -21687,7 +21696,7 @@ LEAN_EXPORT lean_object* l_UInt16_reduceBEq____x40_Lean_Meta_Tactic_Simp_Builtin _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__432; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__433; x_11 = lean_unsigned_to_nat(4u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -21736,7 +21745,7 @@ static lean_object* _init_l___regBuiltin_UInt16_reduceBEq_declare____x40_Lean_Me { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt16_fromExpr___closed__2; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__422; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__423; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -21766,7 +21775,7 @@ static lean_object* _init_l___regBuiltin_UInt16_reduceBEq_declare____x40_Lean_Me { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_UInt16_reduceBEq_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_9210__10188____closed__3; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__57; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -21786,7 +21795,7 @@ static lean_object* _init_l___regBuiltin_UInt16_reduceBEq_declare____x40_Lean_Me { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_UInt16_reduceBEq_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_9210__10188____closed__5; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__59; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -22142,7 +22151,7 @@ LEAN_EXPORT lean_object* l_UInt16_reduceBNe____x40_Lean_Meta_Tactic_Simp_Builtin _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__443; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__444; x_11 = lean_unsigned_to_nat(4u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -22191,7 +22200,7 @@ static lean_object* _init_l___regBuiltin_UInt16_reduceBNe_declare____x40_Lean_Me { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt16_fromExpr___closed__2; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__435; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__436; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -22221,7 +22230,7 @@ static lean_object* _init_l___regBuiltin_UInt16_reduceBNe_declare____x40_Lean_Me { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_UInt16_reduceBNe_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_9210__10228____closed__3; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__57; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -22241,7 +22250,7 @@ static lean_object* _init_l___regBuiltin_UInt16_reduceBNe_declare____x40_Lean_Me { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_UInt16_reduceBNe_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_9210__10228____closed__5; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__59; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -22560,7 +22569,7 @@ static lean_object* _init_l___regBuiltin_UInt16_reduceOfNatCore_declare____x40_L { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt16_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__448; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__449; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -22848,7 +22857,7 @@ static lean_object* _init_l___regBuiltin_UInt16_reduceOfNat_declare____x40_Lean_ { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt16_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__474; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__475; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -22869,7 +22878,7 @@ static lean_object* _init_l___regBuiltin_UInt16_reduceOfNat_declare____x40_Lean_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__27; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__28; x_2 = l___regBuiltin_UInt16_reduceOfNat_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_10455____closed__2; x_3 = lean_array_push(x_1, x_2); return x_3; @@ -23114,7 +23123,7 @@ static lean_object* _init_l___regBuiltin_UInt16_reduceToNat_declare____x40_Lean_ { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt16_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__478; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__479; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -23135,7 +23144,7 @@ static lean_object* _init_l___regBuiltin_UInt16_reduceToNat_declare____x40_Lean_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__27; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__28; x_2 = l___regBuiltin_UInt16_reduceToNat_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_10570____closed__2; x_3 = lean_array_push(x_1, x_2); return x_3; @@ -23208,7 +23217,7 @@ LEAN_EXPORT lean_object* l_UInt16_isValue____x40_Lean_Meta_Tactic_Simp_BuiltinSi _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__494; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__495; x_11 = lean_unsigned_to_nat(3u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -23250,7 +23259,7 @@ static lean_object* _init_l___regBuiltin_UInt16_isValue_declare____x40_Lean_Meta { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt16_fromExpr___closed__2; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__488; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__489; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -23280,7 +23289,7 @@ static lean_object* _init_l___regBuiltin_UInt16_isValue_declare____x40_Lean_Meta { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_UInt16_isValue_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_9210__10643____closed__3; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__57; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -23300,7 +23309,7 @@ static lean_object* _init_l___regBuiltin_UInt16_isValue_declare____x40_Lean_Meta { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_UInt16_isValue_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_9210__10643____closed__5; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__59; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -24572,7 +24581,7 @@ LEAN_EXPORT lean_object* l_UInt32_reduceAdd(lean_object* x_1, lean_object* x_2, _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__277; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__278; x_11 = lean_unsigned_to_nat(6u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -24621,7 +24630,7 @@ static lean_object* _init_l___regBuiltin_UInt32_reduceAdd_declare____x40_Lean_Me { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt32_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__262; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__263; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -24963,7 +24972,7 @@ LEAN_EXPORT lean_object* l_UInt32_reduceMul(lean_object* x_1, lean_object* x_2, _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__297; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__298; x_11 = lean_unsigned_to_nat(6u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -25012,7 +25021,7 @@ static lean_object* _init_l___regBuiltin_UInt32_reduceMul_declare____x40_Lean_Me { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt32_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__287; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__288; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -25342,7 +25351,7 @@ LEAN_EXPORT lean_object* l_UInt32_reduceSub(lean_object* x_1, lean_object* x_2, _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__310; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__311; x_11 = lean_unsigned_to_nat(6u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -25391,7 +25400,7 @@ static lean_object* _init_l___regBuiltin_UInt32_reduceSub_declare____x40_Lean_Me { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt32_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__300; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__301; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -25721,7 +25730,7 @@ LEAN_EXPORT lean_object* l_UInt32_reduceDiv(lean_object* x_1, lean_object* x_2, _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__323; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__324; x_11 = lean_unsigned_to_nat(6u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -25770,7 +25779,7 @@ static lean_object* _init_l___regBuiltin_UInt32_reduceDiv_declare____x40_Lean_Me { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt32_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__313; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__314; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -26100,7 +26109,7 @@ LEAN_EXPORT lean_object* l_UInt32_reduceMod(lean_object* x_1, lean_object* x_2, _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__336; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__337; x_11 = lean_unsigned_to_nat(6u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -26149,7 +26158,7 @@ static lean_object* _init_l___regBuiltin_UInt32_reduceMod_declare____x40_Lean_Me { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt32_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__326; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__327; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -26449,7 +26458,7 @@ LEAN_EXPORT lean_object* l_UInt32_reduceLT(lean_object* x_1, lean_object* x_2, l _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__352; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__353; x_11 = lean_unsigned_to_nat(4u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -26498,7 +26507,7 @@ static lean_object* _init_l___regBuiltin_UInt32_reduceLT_declare____x40_Lean_Met { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt32_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__342; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__343; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -26778,7 +26787,7 @@ LEAN_EXPORT lean_object* l_UInt32_reduceLE(lean_object* x_1, lean_object* x_2, l _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__366; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__367; x_11 = lean_unsigned_to_nat(4u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -26827,7 +26836,7 @@ static lean_object* _init_l___regBuiltin_UInt32_reduceLE_declare____x40_Lean_Met { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt32_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__356; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__357; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -27107,7 +27116,7 @@ LEAN_EXPORT lean_object* l_UInt32_reduceGT(lean_object* x_1, lean_object* x_2, l _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__379; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__380; x_11 = lean_unsigned_to_nat(4u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -27156,7 +27165,7 @@ static lean_object* _init_l___regBuiltin_UInt32_reduceGT_declare____x40_Lean_Met { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt32_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__369; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__370; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -27396,7 +27405,7 @@ LEAN_EXPORT lean_object* l_UInt32_reduceGE(lean_object* x_1, lean_object* x_2, l _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__392; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__393; x_11 = lean_unsigned_to_nat(4u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -27445,7 +27454,7 @@ static lean_object* _init_l___regBuiltin_UInt32_reduceGE_declare____x40_Lean_Met { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt32_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__382; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__383; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -27685,7 +27694,7 @@ LEAN_EXPORT lean_object* l_UInt32_reduceEq____x40_Lean_Meta_Tactic_Simp_BuiltinS _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__403; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__404; x_11 = lean_unsigned_to_nat(3u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -27734,7 +27743,7 @@ static lean_object* _init_l___regBuiltin_UInt32_reduceEq_declare____x40_Lean_Met { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt32_fromExpr___closed__2; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__395; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__396; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -27764,7 +27773,7 @@ static lean_object* _init_l___regBuiltin_UInt32_reduceEq_declare____x40_Lean_Met { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_UInt32_reduceEq_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_10648__11545____closed__3; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__57; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -27784,7 +27793,7 @@ static lean_object* _init_l___regBuiltin_UInt32_reduceEq_declare____x40_Lean_Met { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_UInt32_reduceEq_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_10648__11545____closed__5; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__59; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -28095,7 +28104,7 @@ LEAN_EXPORT lean_object* l_UInt32_reduceNe____x40_Lean_Meta_Tactic_Simp_BuiltinS _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__417; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__418; x_11 = lean_unsigned_to_nat(3u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -28144,7 +28153,7 @@ static lean_object* _init_l___regBuiltin_UInt32_reduceNe_declare____x40_Lean_Met { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt32_fromExpr___closed__2; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__409; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__410; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -28174,7 +28183,7 @@ static lean_object* _init_l___regBuiltin_UInt32_reduceNe_declare____x40_Lean_Met { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_UInt32_reduceNe_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_10648__11585____closed__3; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__57; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -28194,7 +28203,7 @@ static lean_object* _init_l___regBuiltin_UInt32_reduceNe_declare____x40_Lean_Met { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_UInt32_reduceNe_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_10648__11585____closed__5; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__59; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -28540,7 +28549,7 @@ LEAN_EXPORT lean_object* l_UInt32_reduceBEq____x40_Lean_Meta_Tactic_Simp_Builtin _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__432; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__433; x_11 = lean_unsigned_to_nat(4u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -28589,7 +28598,7 @@ static lean_object* _init_l___regBuiltin_UInt32_reduceBEq_declare____x40_Lean_Me { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt32_fromExpr___closed__2; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__422; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__423; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -28619,7 +28628,7 @@ static lean_object* _init_l___regBuiltin_UInt32_reduceBEq_declare____x40_Lean_Me { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_UInt32_reduceBEq_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_10648__11626____closed__3; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__57; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -28639,7 +28648,7 @@ static lean_object* _init_l___regBuiltin_UInt32_reduceBEq_declare____x40_Lean_Me { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_UInt32_reduceBEq_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_10648__11626____closed__5; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__59; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -28995,7 +29004,7 @@ LEAN_EXPORT lean_object* l_UInt32_reduceBNe____x40_Lean_Meta_Tactic_Simp_Builtin _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__443; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__444; x_11 = lean_unsigned_to_nat(4u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -29044,7 +29053,7 @@ static lean_object* _init_l___regBuiltin_UInt32_reduceBNe_declare____x40_Lean_Me { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt32_fromExpr___closed__2; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__435; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__436; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -29074,7 +29083,7 @@ static lean_object* _init_l___regBuiltin_UInt32_reduceBNe_declare____x40_Lean_Me { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_UInt32_reduceBNe_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_10648__11666____closed__3; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__57; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -29094,7 +29103,7 @@ static lean_object* _init_l___regBuiltin_UInt32_reduceBNe_declare____x40_Lean_Me { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_UInt32_reduceBNe_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_10648__11666____closed__5; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__59; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -29413,7 +29422,7 @@ static lean_object* _init_l___regBuiltin_UInt32_reduceOfNatCore_declare____x40_L { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt32_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__448; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__449; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -29701,7 +29710,7 @@ static lean_object* _init_l___regBuiltin_UInt32_reduceOfNat_declare____x40_Lean_ { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt32_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__474; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__475; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -29722,7 +29731,7 @@ static lean_object* _init_l___regBuiltin_UInt32_reduceOfNat_declare____x40_Lean_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__27; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__28; x_2 = l___regBuiltin_UInt32_reduceOfNat_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_11893____closed__2; x_3 = lean_array_push(x_1, x_2); return x_3; @@ -29967,7 +29976,7 @@ static lean_object* _init_l___regBuiltin_UInt32_reduceToNat_declare____x40_Lean_ { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt32_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__478; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__479; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -29988,7 +29997,7 @@ static lean_object* _init_l___regBuiltin_UInt32_reduceToNat_declare____x40_Lean_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__27; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__28; x_2 = l___regBuiltin_UInt32_reduceToNat_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_12008____closed__2; x_3 = lean_array_push(x_1, x_2); return x_3; @@ -30061,7 +30070,7 @@ LEAN_EXPORT lean_object* l_UInt32_isValue____x40_Lean_Meta_Tactic_Simp_BuiltinSi _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__494; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__495; x_11 = lean_unsigned_to_nat(3u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -30103,7 +30112,7 @@ static lean_object* _init_l___regBuiltin_UInt32_isValue_declare____x40_Lean_Meta { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt32_fromExpr___closed__2; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__488; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__489; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -30133,7 +30142,7 @@ static lean_object* _init_l___regBuiltin_UInt32_isValue_declare____x40_Lean_Meta { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_UInt32_isValue_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_10648__12081____closed__3; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__57; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -30153,7 +30162,7 @@ static lean_object* _init_l___regBuiltin_UInt32_isValue_declare____x40_Lean_Meta { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_UInt32_isValue_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_10648__12081____closed__5; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__59; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -31425,7 +31434,7 @@ LEAN_EXPORT lean_object* l_UInt64_reduceAdd(lean_object* x_1, lean_object* x_2, _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__277; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__278; x_11 = lean_unsigned_to_nat(6u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -31474,7 +31483,7 @@ static lean_object* _init_l___regBuiltin_UInt64_reduceAdd_declare____x40_Lean_Me { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt64_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__262; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__263; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -31816,7 +31825,7 @@ LEAN_EXPORT lean_object* l_UInt64_reduceMul(lean_object* x_1, lean_object* x_2, _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__297; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__298; x_11 = lean_unsigned_to_nat(6u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -31865,7 +31874,7 @@ static lean_object* _init_l___regBuiltin_UInt64_reduceMul_declare____x40_Lean_Me { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt64_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__287; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__288; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -32195,7 +32204,7 @@ LEAN_EXPORT lean_object* l_UInt64_reduceSub(lean_object* x_1, lean_object* x_2, _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__310; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__311; x_11 = lean_unsigned_to_nat(6u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -32244,7 +32253,7 @@ static lean_object* _init_l___regBuiltin_UInt64_reduceSub_declare____x40_Lean_Me { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt64_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__300; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__301; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -32574,7 +32583,7 @@ LEAN_EXPORT lean_object* l_UInt64_reduceDiv(lean_object* x_1, lean_object* x_2, _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__323; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__324; x_11 = lean_unsigned_to_nat(6u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -32623,7 +32632,7 @@ static lean_object* _init_l___regBuiltin_UInt64_reduceDiv_declare____x40_Lean_Me { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt64_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__313; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__314; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -32953,7 +32962,7 @@ LEAN_EXPORT lean_object* l_UInt64_reduceMod(lean_object* x_1, lean_object* x_2, _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__336; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__337; x_11 = lean_unsigned_to_nat(6u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -33002,7 +33011,7 @@ static lean_object* _init_l___regBuiltin_UInt64_reduceMod_declare____x40_Lean_Me { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt64_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__326; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__327; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -33302,7 +33311,7 @@ LEAN_EXPORT lean_object* l_UInt64_reduceLT(lean_object* x_1, lean_object* x_2, l _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__352; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__353; x_11 = lean_unsigned_to_nat(4u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -33351,7 +33360,7 @@ static lean_object* _init_l___regBuiltin_UInt64_reduceLT_declare____x40_Lean_Met { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt64_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__342; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__343; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -33631,7 +33640,7 @@ LEAN_EXPORT lean_object* l_UInt64_reduceLE(lean_object* x_1, lean_object* x_2, l _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__366; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__367; x_11 = lean_unsigned_to_nat(4u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -33680,7 +33689,7 @@ static lean_object* _init_l___regBuiltin_UInt64_reduceLE_declare____x40_Lean_Met { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt64_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__356; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__357; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -33960,7 +33969,7 @@ LEAN_EXPORT lean_object* l_UInt64_reduceGT(lean_object* x_1, lean_object* x_2, l _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__379; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__380; x_11 = lean_unsigned_to_nat(4u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -34009,7 +34018,7 @@ static lean_object* _init_l___regBuiltin_UInt64_reduceGT_declare____x40_Lean_Met { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt64_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__369; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__370; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -34249,7 +34258,7 @@ LEAN_EXPORT lean_object* l_UInt64_reduceGE(lean_object* x_1, lean_object* x_2, l _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__392; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__393; x_11 = lean_unsigned_to_nat(4u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -34298,7 +34307,7 @@ static lean_object* _init_l___regBuiltin_UInt64_reduceGE_declare____x40_Lean_Met { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt64_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__382; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__383; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -34538,7 +34547,7 @@ LEAN_EXPORT lean_object* l_UInt64_reduceEq____x40_Lean_Meta_Tactic_Simp_BuiltinS _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__403; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__404; x_11 = lean_unsigned_to_nat(3u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -34587,7 +34596,7 @@ static lean_object* _init_l___regBuiltin_UInt64_reduceEq_declare____x40_Lean_Met { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt64_fromExpr___closed__2; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__395; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__396; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -34617,7 +34626,7 @@ static lean_object* _init_l___regBuiltin_UInt64_reduceEq_declare____x40_Lean_Met { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_UInt64_reduceEq_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_12086__12983____closed__3; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__57; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -34637,7 +34646,7 @@ static lean_object* _init_l___regBuiltin_UInt64_reduceEq_declare____x40_Lean_Met { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_UInt64_reduceEq_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_12086__12983____closed__5; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__59; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -34948,7 +34957,7 @@ LEAN_EXPORT lean_object* l_UInt64_reduceNe____x40_Lean_Meta_Tactic_Simp_BuiltinS _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__417; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__418; x_11 = lean_unsigned_to_nat(3u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -34997,7 +35006,7 @@ static lean_object* _init_l___regBuiltin_UInt64_reduceNe_declare____x40_Lean_Met { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt64_fromExpr___closed__2; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__409; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__410; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -35027,7 +35036,7 @@ static lean_object* _init_l___regBuiltin_UInt64_reduceNe_declare____x40_Lean_Met { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_UInt64_reduceNe_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_12086__13023____closed__3; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__57; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -35047,7 +35056,7 @@ static lean_object* _init_l___regBuiltin_UInt64_reduceNe_declare____x40_Lean_Met { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_UInt64_reduceNe_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_12086__13023____closed__5; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__59; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -35393,7 +35402,7 @@ LEAN_EXPORT lean_object* l_UInt64_reduceBEq____x40_Lean_Meta_Tactic_Simp_Builtin _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__432; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__433; x_11 = lean_unsigned_to_nat(4u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -35442,7 +35451,7 @@ static lean_object* _init_l___regBuiltin_UInt64_reduceBEq_declare____x40_Lean_Me { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt64_fromExpr___closed__2; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__422; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__423; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -35472,7 +35481,7 @@ static lean_object* _init_l___regBuiltin_UInt64_reduceBEq_declare____x40_Lean_Me { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_UInt64_reduceBEq_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_12086__13064____closed__3; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__57; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -35492,7 +35501,7 @@ static lean_object* _init_l___regBuiltin_UInt64_reduceBEq_declare____x40_Lean_Me { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_UInt64_reduceBEq_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_12086__13064____closed__5; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__59; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -35848,7 +35857,7 @@ LEAN_EXPORT lean_object* l_UInt64_reduceBNe____x40_Lean_Meta_Tactic_Simp_Builtin _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__443; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__444; x_11 = lean_unsigned_to_nat(4u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -35897,7 +35906,7 @@ static lean_object* _init_l___regBuiltin_UInt64_reduceBNe_declare____x40_Lean_Me { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt64_fromExpr___closed__2; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__435; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__436; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -35927,7 +35936,7 @@ static lean_object* _init_l___regBuiltin_UInt64_reduceBNe_declare____x40_Lean_Me { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_UInt64_reduceBNe_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_12086__13104____closed__3; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__57; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -35947,7 +35956,7 @@ static lean_object* _init_l___regBuiltin_UInt64_reduceBNe_declare____x40_Lean_Me { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_UInt64_reduceBNe_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_12086__13104____closed__5; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__59; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -36266,7 +36275,7 @@ static lean_object* _init_l___regBuiltin_UInt64_reduceOfNatCore_declare____x40_L { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt64_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__448; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__449; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -36554,7 +36563,7 @@ static lean_object* _init_l___regBuiltin_UInt64_reduceOfNat_declare____x40_Lean_ { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt64_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__474; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__475; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -36575,7 +36584,7 @@ static lean_object* _init_l___regBuiltin_UInt64_reduceOfNat_declare____x40_Lean_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__27; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__28; x_2 = l___regBuiltin_UInt64_reduceOfNat_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_13331____closed__2; x_3 = lean_array_push(x_1, x_2); return x_3; @@ -36820,7 +36829,7 @@ static lean_object* _init_l___regBuiltin_UInt64_reduceToNat_declare____x40_Lean_ { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt64_fromExpr___closed__1; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__478; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__479; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -36841,7 +36850,7 @@ static lean_object* _init_l___regBuiltin_UInt64_reduceToNat_declare____x40_Lean_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__27; +x_1 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__28; x_2 = l___regBuiltin_UInt64_reduceToNat_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_13446____closed__2; x_3 = lean_array_push(x_1, x_2); return x_3; @@ -36914,7 +36923,7 @@ LEAN_EXPORT lean_object* l_UInt64_isValue____x40_Lean_Meta_Tactic_Simp_BuiltinSi _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__494; +x_10 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__495; x_11 = lean_unsigned_to_nat(3u); x_12 = l_Lean_Expr_isAppOfArity(x_1, x_10, x_11); if (x_12 == 0) @@ -36956,7 +36965,7 @@ static lean_object* _init_l___regBuiltin_UInt64_isValue_declare____x40_Lean_Meta { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_UInt64_fromExpr___closed__2; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__488; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__489; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -36986,7 +36995,7 @@ static lean_object* _init_l___regBuiltin_UInt64_isValue_declare____x40_Lean_Meta { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_UInt64_isValue_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_12086__13519____closed__3; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__57; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -37006,7 +37015,7 @@ static lean_object* _init_l___regBuiltin_UInt64_isValue_declare____x40_Lean_Meta { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_UInt64_isValue_declare____x40_Lean_Meta_Tactic_Simp_BuiltinSimprocs_UInt___hyg_12086__13519____closed__5; -x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__58; +x_2 = l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__59; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -38163,6 +38172,8 @@ l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__comman lean_mark_persistent(l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__501); l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__502 = _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__502(); lean_mark_persistent(l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__502); +l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__503 = _init_l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__503(); +lean_mark_persistent(l___aux__Lean__Meta__Tactic__Simp__BuiltinSimprocs__UInt______macroRules__commandDeclare__uint__simprocs____1___closed__503); l_UInt8_fromExpr___closed__1 = _init_l_UInt8_fromExpr___closed__1(); lean_mark_persistent(l_UInt8_fromExpr___closed__1); l_UInt8_fromExpr___closed__2 = _init_l_UInt8_fromExpr___closed__2(); diff --git a/stage0/stdlib/Lean/Parser/Command.c b/stage0/stdlib/Lean/Parser/Command.c index 354f895cc9dc..fbd54b9d4ffe 100644 --- a/stage0/stdlib/Lean/Parser/Command.c +++ b/stage0/stdlib/Lean/Parser/Command.c @@ -30,7 +30,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_initialize_parenthesizer(lean_obj static lean_object* l_Lean_Parser_Command_openRenaming___closed__6; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_openHiding_parenthesizer(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Command_mutual; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__3; static lean_object* l_Lean_Parser_Command_declaration_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Command_declId_parenthesizer___closed__10; lean_object* l_Lean_PrettyPrinter_Parenthesizer_recover_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -56,8 +55,8 @@ static lean_object* l_Lean_Parser_Command_declSig_formatter___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_end_declRange___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Command_check_formatter___closed__1; static lean_object* l_Lean_Parser_Command_deriving_parenthesizer___closed__8; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__34; LEAN_EXPORT lean_object* l_Lean_Parser_Command_extends_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__44; static lean_object* l_Lean_Parser_Command_example___closed__4; static lean_object* l_Lean_Parser_Command_namedPrio_parenthesizer___closed__11; static lean_object* l_Lean_Parser_Command_check_formatter___closed__3; @@ -69,6 +68,7 @@ static lean_object* l_Lean_Parser_Command_exit_parenthesizer___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_print(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Command_structImplicitBinder; static lean_object* l_Lean_Parser_Command_deriveInduction___closed__4; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__14; static lean_object* l_Lean_Parser_Command_printEqns_formatter___closed__1; static lean_object* l_Lean_Parser_Command_inductive___closed__21; static lean_object* l_Lean_Parser_Command_end___closed__4; @@ -113,7 +113,6 @@ static lean_object* l_Lean_Parser_Command_definition___closed__10; 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*); lean_object* l_Lean_Parser_Term_letDecl_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__29; static lean_object* l_Lean_Parser_Command_declModifiers___closed__12; static lean_object* l_Lean_Parser_Command_genInjectiveTheorems___closed__3; static lean_object* l_Lean_Parser_Command_print___closed__5; @@ -220,7 +219,6 @@ static lean_object* l_Lean_Parser_Tactic_open_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Tactic_open_formatter___closed__8; static lean_object* l_Lean_Parser_Command_structInstBinder___closed__2; 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; static lean_object* l_Lean_Parser_Command_openHiding_formatter___closed__5; static lean_object* l_Lean_Parser_Command_inductive_parenthesizer___closed__4; @@ -228,6 +226,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_variable_declRange___clos static lean_object* l_Lean_Parser_Command_inductive_parenthesizer___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Command_abbrev_formatter___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_private_parenthesizer(lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__30; static lean_object* l_Lean_Parser_Command_declModifiers___closed__7; static lean_object* l_Lean_Parser_Command_declModifiers___closed__14; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_private_formatter(lean_object*); @@ -250,7 +249,6 @@ static lean_object* l_Lean_Parser_Command_eoi_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_openSimple___closed__2; lean_object* l_Lean_Parser_Term_whereDecls_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_example_parenthesizer___closed__1; -static lean_object* l___regBuiltin_Lean_Parser_Command_def_parenthesizer___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Command_section; static lean_object* l___regBuiltin_Lean_Parser_Tactic_open_formatter___closed__1; static lean_object* l_Lean_Parser_Command_addDocString___closed__3; @@ -261,7 +259,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_printAxioms_parenthesizer 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_openOnly_parenthesizer___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__20; 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; @@ -276,6 +273,7 @@ static lean_object* l_Lean_Parser_Command_ctor___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Command_open_declRange___closed__7; static lean_object* l_Lean_Parser_Command_printEqns_formatter___closed__6; static lean_object* l_Lean_Parser_Command_inductive_formatter___closed__10; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__27; 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*); @@ -328,7 +326,6 @@ static lean_object* l_Lean_Parser_Command_classInductive_parenthesizer___closed_ static lean_object* l_Lean_Parser_Command_partial___closed__1; static lean_object* l_Lean_Parser_Command_mutual___closed__7; static lean_object* l_Lean_Parser_Command_optDeclSig___closed__7; -LEAN_EXPORT lean_object* l_Lean_Parser_Command_def_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_quot___closed__2; static lean_object* l_Lean_Parser_Command_structureTk___closed__2; static lean_object* l_Lean_Parser_Command_openSimple_parenthesizer___closed__1; @@ -363,8 +360,6 @@ static lean_object* l_Lean_Parser_Command_whereStructInst___closed__20; static lean_object* l___regBuiltin_Lean_Parser_Command_eraseAttr_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_declId_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_synth_formatter___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__40; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__14; static lean_object* l_Lean_Parser_Command_declaration_formatter___closed__12; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_universe_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Command_theorem___closed__9; @@ -424,7 +419,6 @@ static lean_object* l_Lean_Parser_Command_declaration___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Command_universe_declRange___closed__4; lean_object* l_Lean_Parser_Term_matchAltsWhereDecls_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_structureTk___closed__4; -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_def_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Command_classInductive___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Term_precheckedQuot_declRange___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Command_private_formatter___closed__2; @@ -469,7 +463,6 @@ static lean_object* l_Lean_Parser_Command_initialize_parenthesizer___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Command_noncomputableSection_declRange___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Command_declValSimple_formatter___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Command_ctor; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__26; 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_object* l_Lean_Parser_group_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -505,6 +498,7 @@ static lean_object* l_Lean_Parser_Term_set__option___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_moduleDoc_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Command_openOnly_formatter___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_classTk_formatter(lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__18; static lean_object* l_Lean_Parser_Command_eraseAttr___closed__5; static lean_object* l_Lean_Parser_Command_initializeKeyword___closed__11; static lean_object* l___regBuiltin_Lean_Parser_Command_check__failure_declRange___closed__7; @@ -540,7 +534,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Parser_Command_0__Lean_Parser_Command_ static lean_object* l_Lean_Parser_Command_declValEqns_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_structSimpleBinder_formatter___closed__1; static lean_object* l_Lean_Parser_Command_print___closed__1; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__48; static lean_object* l___regBuiltin_Lean_Parser_Command_in_declRange___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Command_variable_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Command_initializeKeyword_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -589,7 +582,6 @@ static lean_object* l_Lean_Parser_Command_moduleDoc_formatter___closed__7; static lean_object* l_Lean_Parser_Command_initializeKeyword_formatter___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Term_open_declRange___closed__1; static lean_object* l_Lean_Parser_Command_instance_formatter___closed__3; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__38; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_eval_declRange(lean_object*); static lean_object* l_Lean_Parser_Command_end___closed__5; static lean_object* l_Lean_Parser_Command_initialize_parenthesizer___closed__10; @@ -602,6 +594,7 @@ 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_computedField___closed__1; static lean_object* l_Lean_Parser_Command_nonrec___closed__6; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__46; 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; @@ -626,6 +619,7 @@ static lean_object* l_Lean_Parser_Command_declValEqns_formatter___closed__3; static lean_object* l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_init__quot___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Command_attribute_declRange___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__28; static lean_object* l_Lean_Parser_Command_printEqns_formatter___closed__2; static lean_object* l_Lean_Parser_Command_section___closed__1; lean_object* l_Lean_Parser_ppDedent_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -637,7 +631,6 @@ static lean_object* l_Lean_Parser_Command_mutual_formatter___closed__7; static lean_object* l_Lean_Parser_Command_section___closed__2; static lean_object* l_Lean_Parser_Command_declValSimple_formatter___closed__6; static lean_object* l_Lean_Parser_Command_structure___closed__12; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__45; static lean_object* l_Lean_Parser_Command_noncomputableSection___closed__3; static lean_object* l_Lean_Parser_Command_structImplicitBinder_parenthesizer___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Command_structureTk_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -655,6 +648,7 @@ static lean_object* l_Lean_Parser_Command_classInductive___closed__11; static lean_object* l_Lean_Parser_Command_declModifiers_formatter___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Command_computedFields_parenthesizer___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_declValSimple_parenthesizer(lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__29; static lean_object* l_Lean_Parser_Command_genInjectiveTheorems___closed__6; static lean_object* l_Lean_Parser_Command_instance___closed__1; static lean_object* l_Lean_Parser_Command_whereStructInst_formatter___closed__3; @@ -710,7 +704,6 @@ lean_object* l_Lean_Parser_incQuotDepth_formatter(lean_object*, lean_object*, le static lean_object* l___regBuiltin_Lean_Parser_Command_deriving_declRange___closed__2; static lean_object* l_Lean_Parser_Command_classInductive_parenthesizer___closed__8; extern lean_object* l_Lean_Parser_Term_matchAltsWhereDecls; -LEAN_EXPORT lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2829_(lean_object*); static lean_object* l_Lean_Parser_Command_structExplicitBinder___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Command_classInductive; static lean_object* l_Lean_Parser_Command_structImplicitBinder_formatter___closed__1; @@ -720,7 +713,6 @@ static lean_object* l_Lean_Parser_Command_optDefDeriving___closed__2; static lean_object* l_Lean_Parser_Tactic_set__option_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Command_mutual___closed__9; static lean_object* l_Lean_Parser_Command_declValSimple_parenthesizer___closed__10; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__31; static lean_object* l_Lean_Parser_Command_axiom_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_noncomputableSection___closed__10; static lean_object* l___regBuiltin_Lean_Parser_Command_check__failure_declRange___closed__5; @@ -783,6 +775,7 @@ static lean_object* l_Lean_Parser_Command_nonrec_formatter___closed__1; static lean_object* l_Lean_Parser_Command_noncomputableSection___closed__4; static lean_object* l_Lean_Parser_Command_optDeriving___closed__9; LEAN_EXPORT lean_object* l_Lean_Parser_Command_mutual_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__35; static lean_object* l_Lean_Parser_Command_optDeriving_formatter___closed__2; static lean_object* l_Lean_Parser_Command_quot_parenthesizer___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Command_classTk_parenthesizer___closed__2; @@ -873,7 +866,6 @@ static lean_object* l_Lean_Parser_Command_inductive_formatter___closed__15; static lean_object* l_Lean_Parser_Command_structExplicitBinder___closed__8; static lean_object* l_Lean_Parser_Command_optDeriving_formatter___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Command_declId_formatter___closed__2; -static lean_object* l_Lean_Parser_Command_def_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_initialize_formatter___closed__11; static lean_object* l_Lean_Parser_Command_variable___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_printAxioms_declRange(lean_object*); @@ -883,7 +875,6 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_variable_declRange(l static lean_object* l_Lean_Parser_Command_structFields___closed__8; lean_object* l_Lean_Parser_Term_attrInstance_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_Term_matchAltsWhereDecls_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__50; 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; @@ -899,6 +890,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_set__option_parenthesizer static lean_object* l_Lean_Parser_Command_structFields_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_exit___closed__7; static lean_object* l_Lean_Parser_Command_check__failure___closed__5; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__37; static lean_object* l_Lean_Parser_Command_initializeKeyword_formatter___closed__2; static lean_object* l_Lean_Parser_Command_declId_formatter___closed__12; static lean_object* l_Lean_Parser_Command_eoi___closed__5; @@ -923,7 +915,6 @@ 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_2851____closed__19; 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; @@ -939,7 +930,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_section_declRange___close static lean_object* l___regBuiltin_Lean_Parser_Command_deriving_declRange___closed__5; static lean_object* l_Lean_Parser_Command_check_parenthesizer___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Command_deriving_formatter___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__44; 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; @@ -966,6 +956,7 @@ static lean_object* l_Lean_Parser_Command_initialize___closed__9; 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_initFn____x40_Lean_Parser_Command___hyg_2813____closed__12; 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_printEqns_parenthesizer___closed__5; @@ -1045,6 +1036,7 @@ static lean_object* l_Lean_Parser_Command_structInstBinder___closed__14; static lean_object* l_Lean_Parser_Command_eval___closed__4; static lean_object* l_Lean_Parser_Command_printEqns_formatter___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Command_structInstBinder_parenthesizer___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__8; static lean_object* l_Lean_Parser_Command_attribute_formatter___closed__9; static lean_object* l_Lean_Parser_Command_definition_formatter___closed__9; LEAN_EXPORT lean_object* l_Lean_Parser_Command_genInjectiveTheorems_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1131,7 +1123,6 @@ static lean_object* l_Lean_Parser_Command_computedField_parenthesizer___closed__ static lean_object* l_Lean_Parser_Command_optDeriving_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_openSimple_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_whereStructInst_formatter___closed__1; -static lean_object* l_Lean_Parser_Command_def___closed__4; static lean_object* l_Lean_Parser_Command_declId___closed__13; static lean_object* l_Lean_Parser_Command_structSimpleBinder_formatter___closed__6; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_mutual_parenthesizer(lean_object*); @@ -1200,7 +1191,7 @@ static lean_object* l_Lean_Parser_Command_optDefDeriving___closed__3; static lean_object* l_Lean_Parser_Command_declModifiers_formatter___closed__22; static lean_object* l_Lean_Parser_Command_axiom_parenthesizer___closed__4; lean_object* l_Lean_Parser_Term_binderIdent_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_def_parenthesizer___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_declValEqns_formatter(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_definition_parenthesizer___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_in_parenthesizer(lean_object*); @@ -1215,10 +1206,8 @@ static lean_object* l_Lean_Parser_Command_end_parenthesizer___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Command_structFields; static lean_object* l_Lean_Parser_Command_in_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_openScoped_parenthesizer___closed__1; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__8; static lean_object* l_Lean_Parser_Command_genInjectiveTheorems_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Command_check__failure___closed__8; -static lean_object* l_Lean_Parser_Command_declaration_formatter___closed__14; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_variable(lean_object*); static lean_object* l_Lean_Parser_Command_classInductive_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Term_set__option___closed__8; @@ -1273,12 +1262,14 @@ static lean_object* l_Lean_Parser_Command_whereStructInst_parenthesizer___closed static lean_object* l_Lean_Parser_Command_openHiding___closed__7; 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_2813____closed__21; 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; 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_structCtor_parenthesizer___closed__4; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__23; static lean_object* l___regBuiltin_Lean_Parser_Command_export_declRange___closed__4; static lean_object* l_Lean_Parser_Command_openDecl___closed__6; static lean_object* l_Lean_Parser_Command_partial___closed__6; @@ -1314,12 +1305,10 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_eval_declRange___closed__ 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_def_formatter___closed__1; static lean_object* l_Lean_Parser_Command_openOnly___closed__8; static lean_object* l_Lean_Parser_Command_extends___closed__8; static lean_object* l_Lean_Parser_Command_structInstBinder_formatter___closed__5; static lean_object* l_Lean_Parser_Command_classInductive___closed__1; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__53; 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; @@ -1343,7 +1332,6 @@ static lean_object* l_Lean_Parser_Command_declSig___closed__11; static lean_object* l_Lean_Parser_Command_ctor___closed__14; static lean_object* l_Lean_Parser_Command_declValSimple___closed__2; static lean_object* l_Lean_Parser_Command_attribute___closed__2; -LEAN_EXPORT lean_object* l_Lean_Parser_Command_def; static lean_object* l_Lean_Parser_Command_namedPrio_formatter___closed__1; static lean_object* l_Lean_Parser_Command_opaque___closed__11; static lean_object* l_Lean_Parser_Command_theorem___closed__5; @@ -1367,7 +1355,6 @@ static lean_object* l_Lean_Parser_Command_structInstBinder_parenthesizer___close static lean_object* l_Lean_Parser_Command_namedPrio_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Command_namespace_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_deriveInduction_formatter___closed__4; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__46; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_addDocString_parenthesizer(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_declaration_declRange___closed__7; static lean_object* l_Lean_Parser_Command_declId___closed__3; @@ -1402,6 +1389,7 @@ static lean_object* l_Lean_Parser_Command_reduce___closed__7; static lean_object* l_Lean_Parser_Command_open_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_check__failure_formatter___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Command_derivingClasses; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__53; static lean_object* l_Lean_Parser_Tactic_open_formatter___closed__2; static lean_object* l_Lean_Parser_Command_declId___closed__7; static lean_object* l_Lean_Parser_Command_structImplicitBinder___closed__10; @@ -1413,7 +1401,6 @@ static lean_object* l_Lean_Parser_Command_structImplicitBinder_formatter___close static lean_object* l_Lean_Parser_Command_deriving___closed__2; static lean_object* l_Lean_Parser_Command_declModifiers_formatter___closed__9; static lean_object* l_Lean_Parser_Command_opaque_parenthesizer___closed__4; -static lean_object* l_Lean_Parser_Command_def___closed__1; static lean_object* l_Lean_Parser_Command_open_formatter___closed__4; static lean_object* l_Lean_Parser_Command_structure_formatter___closed__3; static lean_object* l_Lean_Parser_Command_openHiding_formatter___closed__3; @@ -1428,12 +1415,10 @@ static lean_object* l_Lean_Parser_Command_openScoped___closed__6; static lean_object* l_Lean_Parser_Command_computedFields___closed__9; static lean_object* l_Lean_Parser_Command_eraseAttr_formatter___closed__2; static lean_object* l_Lean_Parser_Command_instance_formatter___closed__4; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__12; LEAN_EXPORT lean_object* l_Lean_Parser_Command_abbrev; static lean_object* l_Lean_Parser_Command_whereStructField___closed__4; static lean_object* l_Lean_Parser_Command_whereStructInst_formatter___closed__2; static lean_object* l_Lean_Parser_Command_classInductive_formatter___closed__11; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__52; static lean_object* l_Lean_Parser_Command_optDefDeriving_formatter___closed__1; static lean_object* l_Lean_Parser_Command_whereStructInst___closed__18; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_quot_declRange(lean_object*); @@ -1453,7 +1438,6 @@ static lean_object* l_Lean_Parser_Command_initializeKeyword___closed__10; static lean_object* l_Lean_Parser_Command_whereStructInst___closed__28; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_check_formatter(lean_object*); static lean_object* l_Lean_Parser_Tactic_set__option___closed__4; -static lean_object* l_Lean_Parser_Command_declaration_parenthesizer___closed__14; static lean_object* l_Lean_Parser_Command_openOnly_formatter___closed__2; static lean_object* l_Lean_Parser_Command_openRenamingItem_formatter___closed__1; static lean_object* l_Lean_Parser_Command_quot___closed__7; @@ -1470,6 +1454,7 @@ static lean_object* l___private_Lean_Parser_Command_0__Lean_Parser_Command_skipU LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_in_formatter(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_openSimple_parenthesizer___closed__1; lean_object* l_Lean_Parser_Tactic_tacticSeq_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__45; static lean_object* l___regBuiltin_Lean_Parser_Command_export_declRange___closed__3; static lean_object* l_Lean_Parser_Command_openRenamingItem___closed__7; static lean_object* l_Lean_Parser_Command_mutual_formatter___closed__10; @@ -1489,7 +1474,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_eval_parenthesizer___clos LEAN_EXPORT lean_object* l_Lean_Parser_Command_opaque; static lean_object* l_Lean_Parser_Command_declSig___closed__1; static lean_object* l_Lean_Parser_Command_structInstBinder___closed__7; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__36; static lean_object* l___regBuiltin_Lean_Parser_Command_extends_formatter___closed__1; static lean_object* l_Lean_Parser_Command_quot_formatter___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Command_noncomputable_parenthesizer___closed__1; @@ -1522,6 +1506,7 @@ static lean_object* l_Lean_Parser_Command_optDeriving_formatter___closed__3; 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_initFn____x40_Lean_Parser_Command___hyg_2813____closed__51; static lean_object* l_Lean_Parser_Command_variable_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_initialize_parenthesizer___closed__9; static lean_object* l_Lean_Parser_Command_theorem___closed__11; @@ -1535,11 +1520,11 @@ static lean_object* l_Lean_Parser_Command_instance_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_declId___closed__17; 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_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__30; 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_structSimpleBinder_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_declModifiers_parenthesizer___closed__13; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__38; 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; @@ -1547,6 +1532,7 @@ static lean_object* l_Lean_Parser_Command_whereStructInst___closed__14; static lean_object* l_Lean_Parser_Command_whereStructInst_formatter___closed__8; static lean_object* l_Lean_Parser_Command_classInductive_formatter___closed__9; static lean_object* l_Lean_Parser_Command_initialize_parenthesizer___closed__2; +LEAN_EXPORT lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813_(lean_object*); static lean_object* l_Lean_Parser_Command_declaration___closed__12; static lean_object* l_Lean_Parser_Command_computedFields___closed__8; static lean_object* l___regBuiltin_Lean_Parser_Command_mutual_declRange___closed__4; @@ -1580,7 +1566,6 @@ static lean_object* l_Lean_Parser_Command_noncomputableSection_parenthesizer___c static lean_object* l_Lean_Parser_Command_initialize_parenthesizer___closed__12; static lean_object* l___regBuiltin_Lean_Parser_Command_reduce_declRange___closed__2; static lean_object* l_Lean_Parser_Command_variable_formatter___closed__1; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__9; 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; @@ -1599,6 +1584,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_exit_formatter(lean_ static lean_object* l___regBuiltin_Lean_Parser_Command_check__failure_declRange___closed__2; static lean_object* l_Lean_Parser_Command_reduce___closed__4; static lean_object* l_Lean_Parser_Command_structCtor___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__40; static lean_object* l_Lean_Parser_Command_noncomputable_formatter___closed__3; static lean_object* l_Lean_Parser_Command_exit_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__10; @@ -1636,7 +1622,6 @@ static lean_object* l_Lean_Parser_Command_openScoped___closed__8; 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___regBuiltin_Lean_Parser_Command_eraseAttr_formatter___closed__1; static lean_object* l_Lean_Parser_Command_mutual___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__28; static lean_object* l_Lean_Parser_Command_eraseAttr_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Command_ctor_parenthesizer___closed__5; LEAN_EXPORT uint8_t l_Lean_Parser_Command_declId___lambda__1(uint32_t); @@ -1660,11 +1645,9 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_init__quot_parenthes static lean_object* l_Lean_Parser_Command_abbrev___closed__6; static lean_object* l_Lean_Parser_Command_noncomputable___closed__5; static lean_object* l_Lean_Parser_Command_attribute___closed__7; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__16; lean_object* l_Lean_Parser_takeWhileFn___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_openRenaming_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Command_printEqns___closed__5; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__39; static lean_object* l_Lean_Parser_Command_openRenaming_formatter___closed__6; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_abbrev_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Command_declVal___closed__3; @@ -1782,7 +1765,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_protected_parenthesizer__ 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_declaration___closed__18; 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; @@ -1790,7 +1772,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_attribute_declRange___clo static lean_object* l___regBuiltin_Lean_Parser_Command_check_declRange___closed__7; static lean_object* l_Lean_Parser_Command_noncomputableSection_formatter___closed__1; static lean_object* l_Lean_Parser_Command_inductive___closed__10; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__2; static lean_object* l_Lean_Parser_Command_printAxioms_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_structure_formatter___closed__18; static lean_object* l_Lean_Parser_Command_declSig_parenthesizer___closed__3; @@ -1803,6 +1784,7 @@ 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; 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_initFn____x40_Lean_Parser_Command___hyg_2813____closed__47; 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*); static lean_object* l___regBuiltin_Lean_Parser_Command_set__option_declRange___closed__5; @@ -1856,8 +1838,6 @@ static lean_object* l_Lean_Parser_Command_synth_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_declaration_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_genInjectiveTheorems_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_eoi_parenthesizer___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__5; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__25; 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; @@ -1865,7 +1845,6 @@ 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_inductive_parenthesizer___closed__13; static lean_object* l_Lean_Parser_Command_quot___closed__3; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__54; static lean_object* l_Lean_Parser_Command_partial___closed__7; static lean_object* l_Lean_Parser_Command_structCtor___closed__7; static lean_object* l_Lean_Parser_Command_initialize___closed__2; @@ -1889,6 +1868,7 @@ 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_2813____closed__16; static lean_object* l___private_Lean_Parser_Command_0__Lean_Parser_Command_skipUntilWsOrDelim___closed__1; static lean_object* l_Lean_Parser_Command_synth___closed__9; static lean_object* l_Lean_Parser_Command_optDefDeriving___closed__9; @@ -1943,6 +1923,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_genInjectiveTheorems_decl static lean_object* l_Lean_Parser_Command_definition___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Tactic_open_declRange___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Command_private_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__19; LEAN_EXPORT lean_object* l_Lean_Parser_Command_deriveInduction_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_initialize___closed__5; static lean_object* l_Lean_Parser_Tactic_open_parenthesizer___closed__3; @@ -1962,20 +1943,18 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_structure_formatter___clo static lean_object* l_Lean_Parser_Term_set__option_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Command_structFields___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Command_namedPrio_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___regBuiltin_Lean_Parser_Command_def_formatter___closed__2; static lean_object* l_Lean_Parser_Command_print_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_initialize___closed__11; static lean_object* l_Lean_Parser_Command_noncomputable_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_genInjectiveTheorems_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_optDeriving___closed__4; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__42; static lean_object* l_Lean_Parser_Command_namespace_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Command_namedPrio_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_universe_declRange___closed__7; static lean_object* l_Lean_Parser_Command_definition_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_structure_parenthesizer___closed__10; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__33; static lean_object* l_Lean_Parser_Command_opaque_parenthesizer___closed__6; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__18; static lean_object* l_Lean_Parser_Term_set__option_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Command_definition; static lean_object* l_Lean_Parser_Command_declModifiers_parenthesizer___closed__12; @@ -2007,13 +1986,13 @@ static lean_object* l_Lean_Parser_Command_definition_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Command_ctor___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Command_axiom_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_axiom___closed__9; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__43; static lean_object* l_Lean_Parser_Command_initialize_formatter___closed__1; static lean_object* l_Lean_Parser_Command_optDefDeriving_parenthesizer___closed__3; lean_object* l_Lean_Parser_Term_ident_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_quot___closed__4; static lean_object* l_Lean_Parser_Command_initialize___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Command_optNamedPrio; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__32; static lean_object* l_Lean_Parser_Command_structImplicitBinder___closed__9; static lean_object* l_Lean_Parser_Command_synth_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_exit_parenthesizer___closed__2; @@ -2027,13 +2006,11 @@ 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; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_eoi_formatter(lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__35; static lean_object* l_Lean_Parser_Command_export___closed__8; static lean_object* l___regBuiltin_Lean_Parser_Command_print_formatter___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Command_export_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_ctor_parenthesizer___closed__9; static lean_object* l_Lean_Parser_Command_declVal_parenthesizer___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__24; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_open_declRange(lean_object*); static lean_object* l_Lean_Parser_Term_precheckedQuot_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_declId_formatter___closed__7; @@ -2063,7 +2040,6 @@ static lean_object* l_Lean_Parser_Command_structSimpleBinder___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_namedPrio_formatter(lean_object*); static lean_object* l_Lean_Parser_Command_opaque___closed__8; static lean_object* l_Lean_Parser_Command_inductive___closed__12; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Command_computedFields_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_inductive___closed__18; static lean_object* l_Lean_Parser_Command_noncomputable_formatter___closed__2; @@ -2071,9 +2047,9 @@ static lean_object* l_Lean_Parser_Command_declValSimple___closed__3; static lean_object* l_Lean_Parser_Command_openHiding_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Command_openOnly_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_List_elem___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__41; static lean_object* l_Lean_Parser_Command_addDocString___closed__9; static lean_object* l_Lean_Parser_Command_initializeKeyword___closed__6; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__50; static lean_object* l_Lean_Parser_Command_theorem_formatter___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Command_noncomputable_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_precheckedQuot_declRange___closed__1; @@ -2090,7 +2066,6 @@ 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_structImplicitBinder___closed__1; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__22; static lean_object* l_Lean_Parser_Command_init__quot___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Command_unsafe; static lean_object* l___regBuiltin_Lean_Parser_Command_set__option_declRange___closed__3; @@ -2126,7 +2101,6 @@ 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; -static lean_object* l_Lean_Parser_Command_def_formatter___closed__2; static lean_object* l_Lean_Parser_Command_openHiding___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Command_classInductive_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Command_declId___lambda__1___closed__3___boxed__const__1; @@ -2166,7 +2140,6 @@ static lean_object* l_Lean_Parser_Command_moduleDoc___closed__1; static lean_object* l_Lean_Parser_Command_structInstBinder_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_whereStructInst___closed__12; static lean_object* l_Lean_Parser_Command_structure_formatter___closed__16; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__15; static lean_object* l_Lean_Parser_Command_opaque___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Command_noncomputable; static lean_object* l___regBuiltin_Lean_Parser_Term_open_declRange___closed__4; @@ -2182,7 +2155,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_universe; static lean_object* l___regBuiltin_Lean_Parser_Command_exit_declRange___closed__7; static lean_object* l_Lean_Parser_Command_print_parenthesizer___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Command_check__failure_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__7; static lean_object* l_Lean_Parser_Command_declaration___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Command_declValSimple_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_structExplicitBinder_formatter___closed__3; @@ -2214,6 +2186,7 @@ static lean_object* l_Lean_Parser_Command_structExplicitBinder_formatter___close static lean_object* l_Lean_Parser_Command_check___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Command_ctor_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_partial_formatter___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__36; static lean_object* l___regBuiltin_Lean_Parser_Command_set__option_declRange___closed__2; static lean_object* l_Lean_Parser_Command_private___closed__3; lean_object* l_Lean_Parser_ppSpace_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -2255,6 +2228,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_mutual_declRange(lea LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_deriveInduction_formatter(lean_object*); static lean_object* l_Lean_Parser_Command_whereStructInst___closed__3; static lean_object* l_Lean_Parser_Command_export___closed__6; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__31; extern lean_object* l_Lean_PrettyPrinter_Formatter_formatterAliasesRef; static lean_object* l___regBuiltin_Lean_Parser_Tactic_open_declRange___closed__5; static lean_object* l_Lean_Parser_Command_unsafe_formatter___closed__3; @@ -2262,10 +2236,10 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_deriveInduction_declRange static lean_object* l___regBuiltin_Lean_Parser_Command_export_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Tactic_open_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_ctor_formatter___closed__1; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__21; static lean_object* l___regBuiltin_Lean_Parser_Command_reduce_declRange___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Term_quot_parenthesizer___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Command_theorem_formatter___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__20; LEAN_EXPORT lean_object* l_Lean_Parser_Command_in_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_structure___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_openScoped_parenthesizer(lean_object*); @@ -2275,7 +2249,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_check_declRange___closed_ 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_open; -static lean_object* l_Lean_Parser_Command_def___closed__2; static lean_object* l_Lean_Parser_Command_computedFields_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_computedField_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_structImplicitBinder_parenthesizer___closed__2; @@ -2285,8 +2258,10 @@ static lean_object* l_Lean_Parser_Command_reduce___closed__1; static lean_object* l_Lean_Parser_Command_computedFields_parenthesizer___closed__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_initFn____x40_Lean_Parser_Command___hyg_2813____closed__41; 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_2813____closed__39; 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; @@ -2312,8 +2287,10 @@ static lean_object* l_Lean_Parser_Command_optDeriving___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_moduleDoc_declRange___closed__3; static lean_object* l_Lean_Parser_Command_whereStructInst___closed__6; static lean_object* l_Lean_Parser_Command_eraseAttr_formatter___closed__3; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__9; static lean_object* l_Lean_Parser_Command_openHiding_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Command_check__failure_formatter___closed__3; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__49; lean_object* l_Lean_Parser_many(lean_object*); 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*); @@ -2361,7 +2338,6 @@ static lean_object* l_Lean_Parser_Command_check__failure___closed__4; 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; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__27; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_import_formatter(lean_object*); static lean_object* l_Lean_Parser_Command_structFields_formatter___closed__7; static lean_object* l_Lean_Parser_Command_unsafe___closed__4; @@ -2371,7 +2347,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_nonrec; static lean_object* l_Lean_Parser_Command_declValEqns___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_initialize_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Command_export_parenthesizer___closed__2; -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_def_formatter(lean_object*); static lean_object* l_Lean_Parser_Command_declValSimple_parenthesizer___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Command_whereStructField_formatter___closed__2; static lean_object* l_Lean_Parser_Command_eval___closed__2; @@ -2410,9 +2385,11 @@ static lean_object* l_Lean_Parser_Command_optDeclSig___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Command_theorem_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_structureTk___closed__5; static lean_object* l_Lean_Parser_Tactic_open_formatter___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__6; lean_object* l_Lean_Parser_Termination_suffix_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_openScoped___closed__3; static lean_object* l_Lean_Parser_Command_moduleDoc_parenthesizer___closed__2; +LEAN_EXPORT lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2791_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Command_whereStructField_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_printAxioms_declRange___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Term_open_declRange___closed__2; @@ -2463,7 +2440,6 @@ static lean_object* l_Lean_Parser_Command_import___closed__4; static lean_object* l_Lean_Parser_Command_open___closed__6; lean_object* l_Lean_Parser_unicodeSymbol_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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___regBuiltin_Lean_Parser_Command_deriveInduction_docString___closed__1; static lean_object* l_Lean_Parser_Command_declId_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_printAxioms_formatter___closed__1; @@ -2502,6 +2478,7 @@ static lean_object* l_Lean_Parser_Command_eval_formatter___closed__4; static lean_object* l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_initialize_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Command_structFields_parenthesizer___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__24; static lean_object* l_Lean_Parser_Command_declSig_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Command_unsafe_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_genInjectiveTheorems___closed__1; @@ -2557,6 +2534,7 @@ static lean_object* l_Lean_Parser_Command_openDecl___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_printAxioms_declRange___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Command_namespace_declRange___closed__4; static lean_object* l_Lean_Parser_Command_definition___closed__12; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__48; static lean_object* l___regBuiltin_Lean_Parser_Command_theorem_formatter___closed__1; static lean_object* l_Lean_Parser_Command_definition___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Command_structSimpleBinder_parenthesizer___closed__2; @@ -2577,6 +2555,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_structureTk_formatter(lean_object static lean_object* l_Lean_Parser_Command_structInstBinder_formatter___closed__1; static lean_object* l_Lean_Parser_Command_structure_formatter___closed__8; static lean_object* l_Lean_Parser_Command_deriveInduction_parenthesizer___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__17; static lean_object* l_Lean_Parser_Command_declValEqns_formatter___closed__1; static lean_object* l_Lean_Parser_Command_derivingClasses_formatter___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Command_openDecl_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -2590,7 +2569,6 @@ static lean_object* l_Lean_Parser_Command_openDecl___closed__8; lean_object* l_Lean_Parser_withCache(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_mutual_formatter___closed__4; static lean_object* l_Lean_Parser_Command_structure___closed__14; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_openRenaming_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Command_optionValue_formatter___closed__4; static lean_object* l_Lean_Parser_Command_theorem_formatter___closed__4; @@ -2693,7 +2671,6 @@ lean_object* l_Lean_Parser_Term_binderDefault_formatter(lean_object*, lean_objec static lean_object* l_Lean_Parser_Command_structure_formatter___closed__19; static lean_object* l___regBuiltin_Lean_Parser_Command_nonrec_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_optDeclSig_formatter___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__11; static lean_object* l_Lean_Parser_Command_whereStructField_formatter___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Command_print_declRange___closed__4; static lean_object* l_Lean_Parser_Term_quot_formatter___closed__5; @@ -2720,7 +2697,6 @@ lean_object* l_Lean_Parser_registerAlias(lean_object*, lean_object*, lean_object static lean_object* l_Lean_Parser_Command_attribute_parenthesizer___closed__8; LEAN_EXPORT lean_object* l_Lean_Parser_Command_optionValue_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Parser_Command_0__Lean_Parser_Command_skipUntil___closed__2; -static lean_object* l_Lean_Parser_Command_def___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Command_moduleDoc_formatter___closed__1; static lean_object* l_Lean_Parser_Command_initializeKeyword_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Command_ctor___closed__12; @@ -2767,7 +2743,6 @@ static lean_object* l_Lean_Parser_Command_openScoped___closed__2; static lean_object* l_Lean_Parser_Command_initialize_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Tactic_set__option_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_quot_formatter___closed__1; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__42; static lean_object* l_Lean_Parser_Command_quot_parenthesizer___closed__3; lean_object* l_Lean_Parser_many1(lean_object*); static lean_object* l_Lean_Parser_Command_example___closed__3; @@ -2813,7 +2788,6 @@ static lean_object* l_Lean_Parser_Command_namedPrio_formatter___closed__12; static lean_object* l_Lean_Parser_Command_printAxioms_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_addDocString_formatter___closed__1; static lean_object* l_Lean_Parser_Command_exit___closed__1; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__6; static lean_object* l_Lean_Parser_Command_quot___closed__9; static lean_object* l___regBuiltin_Lean_Parser_Command_declaration_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_deriving_parenthesizer___closed__3; @@ -2873,12 +2847,14 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_exit_parenthesizer___clos static lean_object* l_Lean_Parser_Command_namedPrio_formatter___closed__8; static lean_object* l_Lean_Parser_Command_moduleDoc_parenthesizer___closed__8; static lean_object* l_Lean_Parser_Command_openHiding_parenthesizer___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__22; LEAN_EXPORT lean_object* l_Lean_Parser_Command_instance_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_ctor___closed__8; static lean_object* l_Lean_Parser_Command_open___closed__9; static lean_object* l___regBuiltin_Lean_Parser_Command_initialize_declRange___closed__1; static lean_object* l_Lean_Parser_Command_export___closed__9; static lean_object* l_Lean_Parser_Command_opaque_parenthesizer___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__26; 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; @@ -2899,7 +2875,6 @@ static lean_object* l_Lean_Parser_Command_declId_formatter___closed__1; lean_object* l_Lean_Parser_Term_leftArrow_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_initialize_declRange___closed__4; static lean_object* l_Lean_Parser_Command_declValSimple___closed__10; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__37; static lean_object* l_Lean_Parser_Command_deriving___closed__9; static lean_object* l_Lean_Parser_Command_opaque_formatter___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Command_declModifiers(uint8_t); @@ -2917,7 +2892,6 @@ static lean_object* l_Lean_Parser_Command_structInstBinder_parenthesizer___close static lean_object* l_Lean_Parser_Command_declaration_formatter___closed__10; static lean_object* l___regBuiltin_Lean_Parser_Command_print_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_structure___closed__7; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__47; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Tactic_open_formatter(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_optDeriving_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_eval___closed__8; @@ -2946,7 +2920,6 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_check__failure_declR static lean_object* l___regBuiltin_Lean_Parser_Command_deriveInduction_declRange___closed__5; static lean_object* l_Lean_Parser_Command_derivingClasses_formatter___closed__5; 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; static lean_object* l_Lean_Parser_Command_moduleDoc_formatter___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_structInstBinder_formatter(lean_object*); @@ -3003,7 +2976,9 @@ static lean_object* l_Lean_Parser_Command_extends_parenthesizer___closed__2; 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; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Command_openRenaming_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__34; static lean_object* l_Lean_Parser_Command_structSimpleBinder___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Command_printEqns_declRange___closed__6; static lean_object* l_Lean_Parser_Command_print___closed__6; @@ -3053,7 +3028,6 @@ static lean_object* l_Lean_Parser_Term_quot___closed__9; static lean_object* l_Lean_Parser_Command_exit_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_inductive_parenthesizer___closed__9; static lean_object* l_Lean_Parser_Command_computedFields_parenthesizer___closed__7; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__13; static lean_object* l_Lean_Parser_Command_ctor_formatter___closed__4; 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; @@ -3078,7 +3052,6 @@ static lean_object* l_Lean_Parser_Command_namedPrio_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Command_opaque_formatter___closed__5; static lean_object* l_Lean_Parser_Command_unsafe___closed__7; static lean_object* l_Lean_Parser_Command_initializeKeyword_formatter___closed__3; -LEAN_EXPORT lean_object* l_Lean_Parser_Command_def_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_declValEqns___closed__2; static lean_object* l_Lean_Parser_Command_whereStructInst___closed__27; static lean_object* l_Lean_Parser_Command_structImplicitBinder_parenthesizer___closed__6; @@ -3128,9 +3101,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_addDocString_formatter(lean_objec static lean_object* l_Lean_Parser_Command_abbrev_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Command_eraseAttr___closed__6; static lean_object* l_Lean_Parser_Command_set__option___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__23; static lean_object* l_Lean_Parser_Command_inductive___closed__9; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__49; static lean_object* l_Lean_Parser_Command_moduleDoc_formatter___closed__6; static lean_object* l_Lean_Parser_Term_open_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_openHiding_parenthesizer___closed__3; @@ -3148,6 +3119,7 @@ static lean_object* l_Lean_Parser_Command_noncomputableSection_formatter___close static lean_object* l_Lean_Parser_Command_declId_formatter___closed__5; static lean_object* l_Lean_Parser_Command_attribute_formatter___closed__5; static lean_object* l_Lean_Parser_Command_openRenamingItem___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__4; static lean_object* l_Lean_Parser_Command_import___closed__6; static lean_object* l_Lean_Parser_Command_abbrev_parenthesizer___closed__8; static lean_object* l_Lean_Parser_Command_noncomputable___closed__8; @@ -3201,6 +3173,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_initialize_parenthesizer_ 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_initFn____x40_Lean_Parser_Command___hyg_2813____closed__33; static lean_object* l_Lean_Parser_Command_structCtor_formatter___closed__7; lean_object* l_instBEq___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_export_formatter___closed__5; @@ -3210,6 +3183,7 @@ static lean_object* l_Lean_Parser_Command_inductive_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_ctor___closed__1; static lean_object* l_Lean_Parser_Command_declModifiers___closed__6; static lean_object* l_Lean_Parser_Command_eraseAttr___closed__8; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__2; static lean_object* l_Lean_Parser_Command_optionValue___closed__4; static lean_object* l_Lean_Parser_Command_mutual___closed__10; LEAN_EXPORT lean_object* l_Lean_Parser_Command_declModifiers_formatter(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -3220,6 +3194,7 @@ static lean_object* l_Lean_Parser_Command_inductive_formatter___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Command_example_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_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_2813____closed__11; static lean_object* l_Lean_Parser_Command_namedPrio___closed__12; static lean_object* l_Lean_Parser_Command_instance___closed__4; static lean_object* l_Lean_Parser_Command_declId___closed__18; @@ -3244,6 +3219,7 @@ static lean_object* l_Lean_Parser_Command_declValSimple_parenthesizer___closed__ static lean_object* l_Lean_Parser_Command_declModifiers_parenthesizer___closed__24; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_abbrev_formatter(lean_object*); static lean_object* l_Lean_Parser_Command_optDeclSig_formatter___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__10; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_instance_formatter(lean_object*); static lean_object* l_Lean_Parser_Term_set__option_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_structCtor_formatter___closed__4; @@ -3375,6 +3351,7 @@ static lean_object* l_Lean_Parser_Command_optDeriving_parenthesizer___closed__4; 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; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__3; static lean_object* l_Lean_Parser_Command_exit_formatter___closed__2; static lean_object* l_Lean_Parser_Command_in_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_abbrev_formatter___closed__8; @@ -3385,11 +3362,11 @@ static lean_object* l_Lean_Parser_Command_quot___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Command_exit_declRange___closed__1; static lean_object* l_Lean_Parser_Command_classInductive___closed__16; static lean_object* l_Lean_Parser_Command_computedFields_parenthesizer___closed__4; -LEAN_EXPORT lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851_(lean_object*); 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_declaration_parenthesizer___closed__13; static lean_object* l_Lean_Parser_Command_openOnly___closed__6; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__43; static lean_object* l_Lean_Parser_Command_structInstBinder_formatter___closed__8; static lean_object* l_Lean_Parser_Command_print_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Command_import_formatter___closed__2; @@ -3404,6 +3381,7 @@ 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_derivingClasses_formatter___closed__1; static lean_object* l_Lean_Parser_Command_declaration___closed__17; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__1; static lean_object* l_Lean_Parser_Command_openRenaming___closed__10; static lean_object* l_Lean_Parser_Command_set__option___closed__1; static lean_object* l_Lean_Parser_Term_quot_formatter___closed__9; @@ -3440,12 +3418,10 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_optDeclSig_formatter___cl static lean_object* l_Lean_Parser_Command_section_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_instance___closed__7; static lean_object* l_Lean_Parser_Command_attribute_formatter___closed__7; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__51; static lean_object* l_Lean_Parser_Command_classTk___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_moduleDoc_declRange___closed__6; static lean_object* l_Lean_Parser_Command_derivingClasses___closed__3; static lean_object* l_Lean_Parser_Command_whereStructField___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__10; static lean_object* l_Lean_Parser_Command_variable_formatter___closed__2; static lean_object* l_Lean_Parser_Command_universe_formatter___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_noncomputableSection(lean_object*); @@ -3457,6 +3433,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_addDocString_parenthesize static lean_object* l_Lean_Parser_Command_attribute___closed__14; static lean_object* l_Lean_Parser_Command_classInductive_parenthesizer___closed__9; static lean_object* l_Lean_Parser_Command_addDocString_formatter___closed__4; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__15; static lean_object* l_Lean_Parser_Command_structSimpleBinder___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Command_ctor_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_declModifiers_formatter___closed__25; @@ -3466,13 +3443,11 @@ static lean_object* l_Lean_Parser_Command_initialize___closed__8; static lean_object* l_Lean_Parser_Command_declaration___closed__4; static lean_object* l_Lean_Parser_Term_quot_formatter___closed__1; static lean_object* l_Lean_Parser_Command_openRenaming_parenthesizer___closed__7; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__32; static lean_object* l_Lean_Parser_Command_whereStructField_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_declModifiers_parenthesizer___closed__25; LEAN_EXPORT lean_object* l_Lean_Parser_Command_structureTk; static lean_object* l_Lean_Parser_Command_mutual___closed__12; static lean_object* l_Lean_Parser_Command_quot_parenthesizer___closed__1; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__17; static lean_object* l_Lean_Parser_Command_structImplicitBinder_parenthesizer___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Command_eoi_formatter___closed__1; static lean_object* l_Lean_Parser_Term_quot_parenthesizer___closed__9; @@ -3488,6 +3463,7 @@ static lean_object* l_Lean_Parser_Command_abbrev_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Command_structure___closed__6; static lean_object* l_Lean_Parser_Command_inductive_formatter___closed__1; static lean_object* l_Lean_Parser_Command_inductive___closed__20; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__52; LEAN_EXPORT lean_object* l_Lean_Parser_Command_namespace; static lean_object* l___regBuiltin_Lean_Parser_Command_print_formatter___closed__1; static lean_object* l_Lean_Parser_Command_mutual_parenthesizer___closed__1; @@ -3501,7 +3477,6 @@ static lean_object* l_Lean_Parser_Command_declModifiers_parenthesizer___closed__ LEAN_EXPORT lean_object* l_Lean_Parser_Command_declValEqns_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_in_declRange___closed__3; 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_synth___closed__8; static lean_object* l_Lean_Parser_Command_variable_formatter___closed__4; static lean_object* l_Lean_Parser_Tactic_open___closed__3; @@ -3533,6 +3508,7 @@ static lean_object* l_Lean_Parser_Command_namespace_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_section_declRange___closed__5; static lean_object* l_Lean_Parser_Command_namespace___closed__5; static lean_object* l_Lean_Parser_Command_openRenamingItem_parenthesizer___closed__5; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__25; static lean_object* l___regBuiltin_Lean_Parser_Command_noncomputable_formatter___closed__1; static lean_object* l_Lean_Parser_Command_inductive_parenthesizer___closed__10; LEAN_EXPORT lean_object* l_Lean_Parser_Command_visibility; @@ -3559,6 +3535,7 @@ static lean_object* l_Lean_Parser_Term_precheckedQuot_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_unsafe_parenthesizer___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Command_open_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_structure_parenthesizer___closed__4; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__54; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_in(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_printAxioms_declRange___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Tactic_set__option_declRange___closed__5; @@ -3604,6 +3581,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_set__option_parenthesizer(lean_objec static lean_object* l_Lean_Parser_Command_instance_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_deriveInduction_declRange___closed__1; static lean_object* l_Lean_Parser_Command_axiom_formatter___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__13; static lean_object* l_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Command_printAxioms_formatter___closed__4; static lean_object* l_Lean_Parser_Command_whereStructInst___closed__9; @@ -8415,77 +8393,6 @@ x_1 = l_Lean_Parser_Command_definition___closed__13; return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_def___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("def", 3); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Command_def___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_def___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_Command_def___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; -x_1 = l_Lean_Parser_Command_def___closed__1; -x_2 = l_Lean_Parser_Command_def___closed__2; -x_3 = 1; -x_4 = 0; -x_5 = l_Lean_Parser_mkAntiquot(x_1, x_2, x_3, x_4); -return x_5; -} -} -static lean_object* _init_l_Lean_Parser_Command_def___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Command_def___closed__2; -x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_definition___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_def___closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_def___closed__3; -x_2 = l_Lean_Parser_Command_def___closed__4; -x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser_Command_def___closed__6() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_def___closed__2; -x_2 = l_Lean_Parser_Command_def___closed__5; -x_3 = l_Lean_Parser_withCache(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser_Command_def() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Parser_Command_def___closed__6; -return x_1; -} -} static lean_object* _init_l_Lean_Parser_Command_theorem___closed__1() { _start: { @@ -11520,7 +11427,7 @@ static lean_object* _init_l_Lean_Parser_Command_declaration___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_def; +x_1 = l_Lean_Parser_Command_definition; x_2 = l_Lean_Parser_Command_declaration___closed__11; x_3 = l_Lean_Parser_orelse(x_1, x_2); return x_3; @@ -11530,7 +11437,7 @@ static lean_object* _init_l_Lean_Parser_Command_declaration___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_definition; +x_1 = l_Lean_Parser_Command_abbrev; x_2 = l_Lean_Parser_Command_declaration___closed__12; x_3 = l_Lean_Parser_orelse(x_1, x_2); return x_3; @@ -11540,49 +11447,39 @@ static lean_object* _init_l_Lean_Parser_Command_declaration___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_abbrev; -x_2 = l_Lean_Parser_Command_declaration___closed__13; -x_3 = l_Lean_Parser_orelse(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser_Command_declaration___closed__15() { -_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_declaration___closed__14; +x_2 = l_Lean_Parser_Command_declaration___closed__13; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_declaration___closed__16() { +static lean_object* _init_l_Lean_Parser_Command_declaration___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_declaration___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_declaration___closed__15; +x_3 = l_Lean_Parser_Command_declaration___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_declaration___closed__17() { +static lean_object* _init_l_Lean_Parser_Command_declaration___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_declaration___closed__3; -x_2 = l_Lean_Parser_Command_declaration___closed__16; +x_2 = l_Lean_Parser_Command_declaration___closed__15; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_declaration___closed__18() { +static lean_object* _init_l_Lean_Parser_Command_declaration___closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_declaration___closed__2; -x_2 = l_Lean_Parser_Command_declaration___closed__17; +x_2 = l_Lean_Parser_Command_declaration___closed__16; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -11591,7 +11488,7 @@ static lean_object* _init_l_Lean_Parser_Command_declaration() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Command_declaration___closed__18; +x_1 = l_Lean_Parser_Command_declaration___closed__17; return x_1; } } @@ -11612,7 +11509,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(201u); +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); @@ -11624,7 +11521,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(204u); +x_1 = lean_unsigned_to_nat(202u); x_2 = lean_unsigned_to_nat(50u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11652,7 +11549,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(201u); +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); @@ -11664,7 +11561,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(201u); +x_1 = lean_unsigned_to_nat(199u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13812,81 +13709,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_Command_def_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_def___closed__1; -x_2 = l_Lean_Parser_Command_def___closed__2; -x_3 = 1; -x_4 = 0; -x_5 = lean_box(x_3); -x_6 = lean_box(x_4); -x_7 = lean_alloc_closure((void*)(l_Lean_Parser_mkAntiquot_formatter___boxed), 9, 4); -lean_closure_set(x_7, 0, x_1); -lean_closure_set(x_7, 1, x_2); -lean_closure_set(x_7, 2, x_5); -lean_closure_set(x_7, 3, x_6); -return x_7; -} -} -static lean_object* _init_l_Lean_Parser_Command_def_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_def___closed__2; -x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_definition_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); -lean_closure_set(x_4, 2, x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_Parser_Command_def_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_Command_def_formatter___closed__1; -x_7 = l_Lean_Parser_Command_def_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___regBuiltin_Lean_Parser_Command_def_formatter___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_1 = l_Lean_Parser_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_def___closed__1; -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; -} -} -static lean_object* _init_l___regBuiltin_Lean_Parser_Command_def_formatter___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Command_def_formatter), 5, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_def_formatter(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_2 = l___regBuiltin_Lean_Parser_Term_quot_formatter___closed__3; -x_3 = l_Lean_Parser_Command_def___closed__2; -x_4 = l___regBuiltin_Lean_Parser_Command_def_formatter___closed__1; -x_5 = l___regBuiltin_Lean_Parser_Command_def_formatter___closed__2; -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_declSig_formatter___closed__1() { _start: { @@ -17435,7 +17257,7 @@ static lean_object* _init_l_Lean_Parser_Command_declaration_formatter___closed__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Parser_Command_def_formatter___closed__2; +x_1 = l___regBuiltin_Lean_Parser_Command_definition_formatter___closed__2; x_2 = l_Lean_Parser_Command_declaration_formatter___closed__9; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_orelse_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -17447,7 +17269,7 @@ static lean_object* _init_l_Lean_Parser_Command_declaration_formatter___closed__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Parser_Command_definition_formatter___closed__2; +x_1 = l___regBuiltin_Lean_Parser_Command_abbrev_formatter___closed__2; x_2 = l_Lean_Parser_Command_declaration_formatter___closed__10; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_orelse_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -17459,33 +17281,21 @@ static lean_object* _init_l_Lean_Parser_Command_declaration_formatter___closed__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Parser_Command_abbrev_formatter___closed__2; -x_2 = l_Lean_Parser_Command_declaration_formatter___closed__11; -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_declaration_formatter___closed__13() { -_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_declaration_formatter___closed__12; +x_2 = l_Lean_Parser_Command_declaration_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_declaration_formatter___closed__14() { +static lean_object* _init_l_Lean_Parser_Command_declaration_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_Command_declaration___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_declaration_formatter___closed__13; +x_3 = l_Lean_Parser_Command_declaration_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); @@ -17498,7 +17308,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_declaration_formatter(lean_object { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_declaration_formatter___closed__1; -x_7 = l_Lean_Parser_Command_declaration_formatter___closed__14; +x_7 = l_Lean_Parser_Command_declaration_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; } @@ -19641,81 +19451,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_Command_def_parenthesizer___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_1 = l_Lean_Parser_Command_def___closed__1; -x_2 = l_Lean_Parser_Command_def___closed__2; -x_3 = 1; -x_4 = 0; -x_5 = lean_box(x_3); -x_6 = lean_box(x_4); -x_7 = lean_alloc_closure((void*)(l_Lean_Parser_mkAntiquot_parenthesizer___boxed), 9, 4); -lean_closure_set(x_7, 0, x_1); -lean_closure_set(x_7, 1, x_2); -lean_closure_set(x_7, 2, x_5); -lean_closure_set(x_7, 3, x_6); -return x_7; -} -} -static lean_object* _init_l_Lean_Parser_Command_def_parenthesizer___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_def___closed__2; -x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_definition_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); -lean_closure_set(x_4, 2, x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_Parser_Command_def_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_Command_def_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Command_def_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___regBuiltin_Lean_Parser_Command_def_parenthesizer___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_1 = l_Lean_Parser_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_def___closed__1; -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; -} -} -static lean_object* _init_l___regBuiltin_Lean_Parser_Command_def_parenthesizer___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Command_def_parenthesizer), 5, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_def_parenthesizer(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_2 = l___regBuiltin_Lean_Parser_Term_quot_parenthesizer___closed__3; -x_3 = l_Lean_Parser_Command_def___closed__2; -x_4 = l___regBuiltin_Lean_Parser_Command_def_parenthesizer___closed__1; -x_5 = l___regBuiltin_Lean_Parser_Command_def_parenthesizer___closed__2; -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_declSig_parenthesizer___closed__1() { _start: { @@ -23266,7 +23001,7 @@ static lean_object* _init_l_Lean_Parser_Command_declaration_parenthesizer___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Parser_Command_def_parenthesizer___closed__2; +x_1 = l___regBuiltin_Lean_Parser_Command_definition_parenthesizer___closed__2; x_2 = l_Lean_Parser_Command_declaration_parenthesizer___closed__9; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_orelse_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -23278,7 +23013,7 @@ static lean_object* _init_l_Lean_Parser_Command_declaration_parenthesizer___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Parser_Command_definition_parenthesizer___closed__2; +x_1 = l___regBuiltin_Lean_Parser_Command_abbrev_parenthesizer___closed__2; x_2 = l_Lean_Parser_Command_declaration_parenthesizer___closed__10; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_orelse_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -23290,33 +23025,21 @@ static lean_object* _init_l_Lean_Parser_Command_declaration_parenthesizer___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Parser_Command_abbrev_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Command_declaration_parenthesizer___closed__11; -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_declaration_parenthesizer___closed__13() { -_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_declaration_parenthesizer___closed__12; +x_2 = l_Lean_Parser_Command_declaration_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_declaration_parenthesizer___closed__14() { +static lean_object* _init_l_Lean_Parser_Command_declaration_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_Command_declaration___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_declaration_parenthesizer___closed__13; +x_3 = l_Lean_Parser_Command_declaration_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); @@ -23329,7 +23052,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_declaration_parenthesizer(lean_ob { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_declaration_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Command_declaration_parenthesizer___closed__14; +x_7 = l_Lean_Parser_Command_declaration_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; } @@ -23551,7 +23274,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(205u); +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); @@ -23563,7 +23286,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(206u); +x_1 = lean_unsigned_to_nat(204u); x_2 = lean_unsigned_to_nat(94u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23591,7 +23314,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(205u); +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); @@ -23603,7 +23326,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(205u); +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); @@ -24079,7 +23802,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(207u); +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); @@ -24091,7 +23814,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(208u); +x_1 = lean_unsigned_to_nat(206u); x_2 = lean_unsigned_to_nat(62u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -24119,7 +23842,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(207u); +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); @@ -24131,7 +23854,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(207u); +x_1 = lean_unsigned_to_nat(205u); x_2 = lean_unsigned_to_nat(50u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -24519,7 +24242,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(209u); +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); @@ -24531,7 +24254,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(210u); +x_1 = lean_unsigned_to_nat(208u); x_2 = lean_unsigned_to_nat(42u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -24559,7 +24282,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(209u); +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); @@ -24571,7 +24294,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(209u); +x_1 = lean_unsigned_to_nat(207u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -24882,7 +24605,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(211u); +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); @@ -24894,7 +24617,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(212u); +x_1 = lean_unsigned_to_nat(210u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -24922,7 +24645,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(211u); +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); @@ -24934,7 +24657,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(211u); +x_1 = lean_unsigned_to_nat(209u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -25281,7 +25004,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(213u); +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); @@ -25293,7 +25016,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(214u); +x_1 = lean_unsigned_to_nat(212u); x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -25321,7 +25044,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(213u); +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); @@ -25333,7 +25056,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(213u); +x_1 = lean_unsigned_to_nat(211u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -25689,7 +25412,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(215u); +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); @@ -25701,7 +25424,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(216u); +x_1 = lean_unsigned_to_nat(214u); x_2 = lean_unsigned_to_nat(55u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -25729,7 +25452,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(215u); +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); @@ -25741,7 +25464,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(215u); +x_1 = lean_unsigned_to_nat(213u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -26117,7 +25840,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(217u); +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); @@ -26129,7 +25852,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(218u); +x_1 = lean_unsigned_to_nat(216u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -26157,7 +25880,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(217u); +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); @@ -26169,7 +25892,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(217u); +x_1 = lean_unsigned_to_nat(215u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -26544,7 +26267,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(219u); +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); @@ -26556,7 +26279,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(220u); +x_1 = lean_unsigned_to_nat(218u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -26584,7 +26307,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(219u); +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); @@ -26596,7 +26319,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(219u); +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); @@ -26951,7 +26674,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(221u); +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); @@ -26963,7 +26686,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(222u); +x_1 = lean_unsigned_to_nat(220u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -26991,7 +26714,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(221u); +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); @@ -27003,7 +26726,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(221u); +x_1 = lean_unsigned_to_nat(219u); x_2 = lean_unsigned_to_nat(43u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27358,7 +27081,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(223u); +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); @@ -27370,7 +27093,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(224u); +x_1 = lean_unsigned_to_nat(222u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27397,7 +27120,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(223u); +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); @@ -27409,7 +27132,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(223u); +x_1 = lean_unsigned_to_nat(221u); x_2 = lean_unsigned_to_nat(36u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27764,7 +27487,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(225u); +x_1 = lean_unsigned_to_nat(223u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27776,7 +27499,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(226u); +x_1 = lean_unsigned_to_nat(224u); x_2 = lean_unsigned_to_nat(24u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27804,7 +27527,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(225u); +x_1 = lean_unsigned_to_nat(223u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27816,7 +27539,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(225u); +x_1 = lean_unsigned_to_nat(223u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28171,7 +27894,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(227u); +x_1 = lean_unsigned_to_nat(225u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28183,7 +27906,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(228u); +x_1 = lean_unsigned_to_nat(226u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28211,7 +27934,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(227u); +x_1 = lean_unsigned_to_nat(225u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28223,7 +27946,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(227u); +x_1 = lean_unsigned_to_nat(225u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28568,7 +28291,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(229u); +x_1 = lean_unsigned_to_nat(227u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28580,7 +28303,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(230u); +x_1 = lean_unsigned_to_nat(228u); x_2 = lean_unsigned_to_nat(9u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28608,7 +28331,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(229u); +x_1 = lean_unsigned_to_nat(227u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28620,7 +28343,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(229u); +x_1 = lean_unsigned_to_nat(227u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28961,7 +28684,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(231u); +x_1 = lean_unsigned_to_nat(229u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28973,7 +28696,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(232u); +x_1 = lean_unsigned_to_nat(230u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29001,7 +28724,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(231u); +x_1 = lean_unsigned_to_nat(229u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29013,7 +28736,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(231u); +x_1 = lean_unsigned_to_nat(229u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29419,7 +29142,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(233u); +x_1 = lean_unsigned_to_nat(231u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29431,7 +29154,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(234u); +x_1 = lean_unsigned_to_nat(232u); x_2 = lean_unsigned_to_nat(51u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29459,7 +29182,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(233u); +x_1 = lean_unsigned_to_nat(231u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29471,7 +29194,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(233u); +x_1 = lean_unsigned_to_nat(231u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29895,7 +29618,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_printEqns_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(233u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29907,7 +29630,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_printEqns_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(234u); x_2 = lean_unsigned_to_nat(86u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29935,7 +29658,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_printEqns_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(233u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29947,7 +29670,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_printEqns_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(233u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30364,7 +30087,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(237u); +x_1 = lean_unsigned_to_nat(235u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30376,7 +30099,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(238u); +x_1 = lean_unsigned_to_nat(236u); x_2 = lean_unsigned_to_nat(13u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30404,7 +30127,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(237u); +x_1 = lean_unsigned_to_nat(235u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30416,7 +30139,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(237u); +x_1 = lean_unsigned_to_nat(235u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30841,7 +30564,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(240u); +x_1 = lean_unsigned_to_nat(238u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30853,7 +30576,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(241u); +x_1 = lean_unsigned_to_nat(239u); x_2 = lean_unsigned_to_nat(50u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30881,7 +30604,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(240u); +x_1 = lean_unsigned_to_nat(238u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30893,7 +30616,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(240u); +x_1 = lean_unsigned_to_nat(238u); x_2 = lean_unsigned_to_nat(42u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -31608,7 +31331,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(244u); +x_1 = lean_unsigned_to_nat(242u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -31620,7 +31343,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(247u); +x_1 = lean_unsigned_to_nat(245u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -31648,7 +31371,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(244u); +x_1 = lean_unsigned_to_nat(242u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -31660,7 +31383,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(244u); +x_1 = lean_unsigned_to_nat(242u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -32405,7 +32128,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(248u); +x_1 = lean_unsigned_to_nat(246u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -32417,7 +32140,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(249u); +x_1 = lean_unsigned_to_nat(247u); x_2 = lean_unsigned_to_nat(50u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -32445,7 +32168,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(248u); +x_1 = lean_unsigned_to_nat(246u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -32457,7 +32180,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(248u); +x_1 = lean_unsigned_to_nat(246u); x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -32866,7 +32589,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(250u); +x_1 = lean_unsigned_to_nat(248u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -32878,7 +32601,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(251u); +x_1 = lean_unsigned_to_nat(249u); x_2 = lean_unsigned_to_nat(10u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -32906,7 +32629,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(250u); +x_1 = lean_unsigned_to_nat(248u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -32918,7 +32641,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(250u); +x_1 = lean_unsigned_to_nat(248u); x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -34047,7 +33770,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(268u); +x_1 = lean_unsigned_to_nat(266u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -34059,7 +33782,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(269u); +x_1 = lean_unsigned_to_nat(267u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -34087,7 +33810,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(268u); +x_1 = lean_unsigned_to_nat(266u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -34099,7 +33822,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(268u); +x_1 = lean_unsigned_to_nat(266u); x_2 = lean_unsigned_to_nat(36u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -36141,7 +35864,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(271u); +x_1 = lean_unsigned_to_nat(269u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -36153,7 +35876,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(273u); +x_1 = lean_unsigned_to_nat(271u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -36181,7 +35904,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(271u); +x_1 = lean_unsigned_to_nat(269u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -36193,7 +35916,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(271u); +x_1 = lean_unsigned_to_nat(269u); x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -36870,7 +36593,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(276u); +x_1 = lean_unsigned_to_nat(274u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -36882,7 +36605,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(278u); +x_1 = lean_unsigned_to_nat(276u); x_2 = lean_unsigned_to_nat(87u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -36910,7 +36633,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(276u); +x_1 = lean_unsigned_to_nat(274u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -36922,7 +36645,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(276u); +x_1 = lean_unsigned_to_nat(274u); x_2 = lean_unsigned_to_nat(42u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -37660,7 +37383,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(280u); +x_1 = lean_unsigned_to_nat(278u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -37672,7 +37395,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(281u); +x_1 = lean_unsigned_to_nat(279u); x_2 = lean_unsigned_to_nat(47u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -37700,7 +37423,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(280u); +x_1 = lean_unsigned_to_nat(278u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -37712,7 +37435,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(280u); +x_1 = lean_unsigned_to_nat(278u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -38057,7 +37780,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_addDocString_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(283u); +x_1 = lean_unsigned_to_nat(281u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -38069,7 +37792,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_addDocString_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(284u); +x_1 = lean_unsigned_to_nat(282u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -38097,7 +37820,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_addDocString_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(283u); +x_1 = lean_unsigned_to_nat(281u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -38109,7 +37832,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_addDocString_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(283u); +x_1 = lean_unsigned_to_nat(281u); x_2 = lean_unsigned_to_nat(42u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -38506,7 +38229,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_deriveInduction_dec _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(294u); +x_1 = lean_unsigned_to_nat(292u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -38518,7 +38241,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_deriveInduction_dec _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(295u); +x_1 = lean_unsigned_to_nat(293u); x_2 = lean_unsigned_to_nat(48u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -38546,7 +38269,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_deriveInduction_dec _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(294u); +x_1 = lean_unsigned_to_nat(292u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -38558,7 +38281,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_deriveInduction_dec _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(294u); +x_1 = lean_unsigned_to_nat(292u); x_2 = lean_unsigned_to_nat(45u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -38931,7 +38654,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_genInjectiveTheorem _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(301u); +x_1 = lean_unsigned_to_nat(299u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -38943,7 +38666,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_genInjectiveTheorem _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(302u); +x_1 = lean_unsigned_to_nat(300u); x_2 = lean_unsigned_to_nat(37u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -38971,7 +38694,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_genInjectiveTheorem _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(301u); +x_1 = lean_unsigned_to_nat(299u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -38983,7 +38706,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_genInjectiveTheorem _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(301u); +x_1 = lean_unsigned_to_nat(299u); x_2 = lean_unsigned_to_nat(50u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -39481,7 +39204,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_2829_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2791_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; @@ -39542,7 +39265,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_2851____closed__1() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -39552,7 +39275,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_2851____closed__2() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__2() { _start: { lean_object* x_1; @@ -39560,19 +39283,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_2851____closed__3() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____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_2851____closed__2; +x_4 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____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_2851____closed__4() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__4() { _start: { lean_object* x_1; lean_object* x_2; @@ -39582,7 +39305,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_2851____closed__5() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__5() { _start: { lean_object* x_1; lean_object* x_2; @@ -39592,7 +39315,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_2851____closed__6() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__6() { _start: { lean_object* x_1; lean_object* x_2; @@ -39602,12 +39325,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_2851____closed__7() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____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_2851____closed__6; +x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__6; x_3 = 1; x_4 = lean_alloc_ctor(0, 2, 1); lean_ctor_set(x_4, 0, x_1); @@ -39616,7 +39339,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_2851____closed__8() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__8() { _start: { lean_object* x_1; @@ -39624,17 +39347,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_2851____closed__9() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__9() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__8; +x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____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_2851____closed__10() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__10() { _start: { lean_object* x_1; @@ -39642,7 +39365,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_2851____closed__11() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__11() { _start: { lean_object* x_1; @@ -39650,17 +39373,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_2851____closed__12() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__12() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__11; +x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____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_2851____closed__13() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__13() { _start: { lean_object* x_1; @@ -39668,7 +39391,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_2851____closed__14() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__14() { _start: { lean_object* x_1; @@ -39676,17 +39399,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_2851____closed__15() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____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_2851____closed__14; +x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____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_2851____closed__16() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__16() { _start: { lean_object* x_1; @@ -39694,19 +39417,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_2851____closed__17() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____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_2851____closed__16; +x_4 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____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_2851____closed__18() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__18() { _start: { lean_object* x_1; lean_object* x_2; @@ -39716,7 +39439,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_2851____closed__19() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__19() { _start: { lean_object* x_1; @@ -39724,17 +39447,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_2851____closed__20() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__20() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__19; +x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____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_2851____closed__21() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__21() { _start: { lean_object* x_1; @@ -39742,17 +39465,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_2851____closed__22() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__22() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__21; +x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____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_2851____closed__23() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -39762,7 +39485,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_2851____closed__24() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__24() { _start: { lean_object* x_1; lean_object* x_2; @@ -39772,7 +39495,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_2851____closed__25() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__25() { _start: { lean_object* x_1; lean_object* x_2; @@ -39782,7 +39505,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_2851____closed__26() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__26() { _start: { lean_object* x_1; lean_object* x_2; @@ -39792,7 +39515,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_2851____closed__27() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__27() { _start: { lean_object* x_1; lean_object* x_2; @@ -39802,7 +39525,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_2851____closed__28() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__28() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -39812,7 +39535,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_2851____closed__29() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__29() { _start: { lean_object* x_1; lean_object* x_2; @@ -39822,7 +39545,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_2851____closed__30() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__30() { _start: { lean_object* x_1; lean_object* x_2; @@ -39832,7 +39555,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_2851____closed__31() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__31() { _start: { lean_object* x_1; lean_object* x_2; @@ -39842,7 +39565,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_2851____closed__32() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__32() { _start: { lean_object* x_1; lean_object* x_2; @@ -39852,7 +39575,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_2851____closed__33() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__33() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -39862,7 +39585,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_2851____closed__34() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__34() { _start: { lean_object* x_1; lean_object* x_2; @@ -39872,7 +39595,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_2851____closed__35() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__35() { _start: { lean_object* x_1; lean_object* x_2; @@ -39882,7 +39605,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_2851____closed__36() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__36() { _start: { lean_object* x_1; lean_object* x_2; @@ -39892,7 +39615,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_2851____closed__37() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__37() { _start: { lean_object* x_1; lean_object* x_2; @@ -39902,7 +39625,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_2851____closed__38() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__38() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -39912,7 +39635,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_2851____closed__39() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__39() { _start: { lean_object* x_1; lean_object* x_2; @@ -39922,7 +39645,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_2851____closed__40() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__40() { _start: { lean_object* x_1; lean_object* x_2; @@ -39932,7 +39655,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_2851____closed__41() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__41() { _start: { lean_object* x_1; lean_object* x_2; @@ -39942,7 +39665,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_2851____closed__42() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__42() { _start: { lean_object* x_1; lean_object* x_2; @@ -39952,7 +39675,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_2851____closed__43() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__43() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -39962,7 +39685,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_2851____closed__44() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__44() { _start: { lean_object* x_1; lean_object* x_2; @@ -39972,7 +39695,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_2851____closed__45() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__45() { _start: { lean_object* x_1; lean_object* x_2; @@ -39982,7 +39705,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_2851____closed__46() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__46() { _start: { lean_object* x_1; lean_object* x_2; @@ -39992,7 +39715,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_2851____closed__47() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__47() { _start: { lean_object* x_1; lean_object* x_2; @@ -40002,7 +39725,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_2851____closed__48() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__48() { _start: { lean_object* x_1; @@ -40010,29 +39733,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_2851____closed__49() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____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_2851____closed__48; +x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____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_2851____closed__50() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____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_2851____closed__48; +x_4 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____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_2851____closed__51() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__51() { _start: { lean_object* x_1; lean_object* x_2; @@ -40042,17 +39765,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_2851____closed__52() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__52() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__50; +x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____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_2851____closed__53() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__53() { _start: { lean_object* x_1; lean_object* x_2; @@ -40062,7 +39785,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_2851____closed__54() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__54() { _start: { lean_object* x_1; lean_object* x_2; @@ -40072,15 +39795,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_2851_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813_(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_2851____closed__1; -x_3 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__3; -x_4 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__4; -x_5 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__5; -x_6 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__7; +x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__1; +x_3 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__3; +x_4 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__4; +x_5 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__5; +x_6 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____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) { @@ -40088,8 +39811,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_2851____closed__10; -x_10 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__9; +x_9 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__10; +x_10 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__9; x_11 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_2, x_10, x_8); if (lean_obj_tag(x_11) == 0) { @@ -40097,8 +39820,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_2851____closed__13; -x_14 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__12; +x_13 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__13; +x_14 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__12; x_15 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_2, x_14, x_12); if (lean_obj_tag(x_15) == 0) { @@ -40106,9 +39829,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_2851____closed__15; -x_18 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__17; -x_19 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__18; +x_17 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__15; +x_18 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__17; +x_19 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____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) { @@ -40116,7 +39839,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_2851____closed__20; +x_22 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__20; x_23 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_17, x_22, x_21); if (lean_obj_tag(x_23) == 0) { @@ -40124,7 +39847,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_2851____closed__22; +x_25 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__22; x_26 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_17, x_25, x_24); if (lean_obj_tag(x_26) == 0) { @@ -40132,10 +39855,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_2851____closed__23; +x_28 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__23; x_29 = l_Lean_Parser_Command_declId___closed__2; -x_30 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__24; -x_31 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__25; +x_30 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__24; +x_31 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____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) { @@ -40143,7 +39866,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_2851____closed__26; +x_34 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__26; x_35 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_28, x_34, x_33); if (lean_obj_tag(x_35) == 0) { @@ -40151,7 +39874,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_2851____closed__27; +x_37 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__27; x_38 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_28, x_37, x_36); if (lean_obj_tag(x_38) == 0) { @@ -40159,10 +39882,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_2851____closed__28; +x_40 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__28; x_41 = l_Lean_Parser_Command_declSig___closed__2; -x_42 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__29; -x_43 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__30; +x_42 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__29; +x_43 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____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) { @@ -40170,7 +39893,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_2851____closed__31; +x_46 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__31; x_47 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_40, x_46, x_45); if (lean_obj_tag(x_47) == 0) { @@ -40178,7 +39901,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_2851____closed__32; +x_49 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__32; x_50 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_40, x_49, x_48); if (lean_obj_tag(x_50) == 0) { @@ -40186,10 +39909,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_2851____closed__33; +x_52 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__33; x_53 = l_Lean_Parser_Command_declVal___closed__2; -x_54 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__34; -x_55 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__35; +x_54 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__34; +x_55 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____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) { @@ -40197,7 +39920,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_2851____closed__36; +x_58 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__36; x_59 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_52, x_58, x_57); if (lean_obj_tag(x_59) == 0) { @@ -40205,7 +39928,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_2851____closed__37; +x_61 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__37; x_62 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_52, x_61, x_60); if (lean_obj_tag(x_62) == 0) { @@ -40213,10 +39936,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_2851____closed__38; +x_64 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__38; x_65 = l_Lean_Parser_Command_optDeclSig___closed__2; -x_66 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__39; -x_67 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__40; +x_66 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__39; +x_67 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____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) { @@ -40224,7 +39947,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_2851____closed__41; +x_70 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__41; x_71 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_64, x_70, x_69); if (lean_obj_tag(x_71) == 0) { @@ -40232,7 +39955,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_2851____closed__42; +x_73 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__42; x_74 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_64, x_73, x_72); if (lean_obj_tag(x_74) == 0) { @@ -40240,10 +39963,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_2851____closed__43; +x_76 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__43; x_77 = l_Lean_Parser_Command_openDecl___closed__2; -x_78 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__44; -x_79 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__45; +x_78 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__44; +x_79 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____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) { @@ -40251,7 +39974,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_2851____closed__46; +x_82 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__46; x_83 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_76, x_82, x_81); if (lean_obj_tag(x_83) == 0) { @@ -40259,7 +39982,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_2851____closed__47; +x_85 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__47; x_86 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_76, x_85, x_84); if (lean_obj_tag(x_86) == 0) { @@ -40267,10 +39990,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_2851____closed__49; -x_89 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__50; -x_90 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__51; -x_91 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__52; +x_88 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__49; +x_89 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__50; +x_90 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__51; +x_91 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____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) { @@ -40278,7 +40001,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_2851____closed__53; +x_94 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__53; x_95 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_88, x_94, x_93); if (lean_obj_tag(x_95) == 0) { @@ -40286,7 +40009,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_2851____closed__54; +x_97 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__54; x_98 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_88, x_97, x_96); return x_98; } @@ -40949,7 +40672,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_open_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(332u); +x_1 = lean_unsigned_to_nat(330u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -40961,7 +40684,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_open_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(333u); +x_1 = lean_unsigned_to_nat(331u); x_2 = lean_unsigned_to_nat(67u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -40989,7 +40712,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_open_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(332u); +x_1 = lean_unsigned_to_nat(330u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -41001,7 +40724,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_open_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(332u); +x_1 = lean_unsigned_to_nat(330u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -41427,7 +41150,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_set__option_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(339u); +x_1 = lean_unsigned_to_nat(337u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -41439,7 +41162,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_set__option_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(340u); +x_1 = lean_unsigned_to_nat(338u); x_2 = lean_unsigned_to_nat(82u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -41467,7 +41190,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_set__option_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(339u); +x_1 = lean_unsigned_to_nat(337u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -41479,7 +41202,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_set__option_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(339u); +x_1 = lean_unsigned_to_nat(337u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -41959,7 +41682,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_open_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(346u); +x_1 = lean_unsigned_to_nat(344u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -41971,7 +41694,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_open_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(347u); +x_1 = lean_unsigned_to_nat(345u); x_2 = lean_unsigned_to_nat(67u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -41999,7 +41722,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_open_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(346u); +x_1 = lean_unsigned_to_nat(344u); x_2 = lean_unsigned_to_nat(29u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -42011,7 +41734,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_open_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(346u); +x_1 = lean_unsigned_to_nat(344u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -42473,7 +42196,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_set__option_declRang _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(351u); +x_1 = lean_unsigned_to_nat(349u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -42485,7 +42208,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_set__option_declRang _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(352u); +x_1 = lean_unsigned_to_nat(350u); x_2 = lean_unsigned_to_nat(81u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -42513,7 +42236,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_set__option_declRang _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(351u); +x_1 = lean_unsigned_to_nat(349u); x_2 = lean_unsigned_to_nat(29u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -42525,7 +42248,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_set__option_declRang _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(351u); +x_1 = lean_unsigned_to_nat(349u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -43730,20 +43453,6 @@ l_Lean_Parser_Command_definition___closed__13 = _init_l_Lean_Parser_Command_defi lean_mark_persistent(l_Lean_Parser_Command_definition___closed__13); l_Lean_Parser_Command_definition = _init_l_Lean_Parser_Command_definition(); lean_mark_persistent(l_Lean_Parser_Command_definition); -l_Lean_Parser_Command_def___closed__1 = _init_l_Lean_Parser_Command_def___closed__1(); -lean_mark_persistent(l_Lean_Parser_Command_def___closed__1); -l_Lean_Parser_Command_def___closed__2 = _init_l_Lean_Parser_Command_def___closed__2(); -lean_mark_persistent(l_Lean_Parser_Command_def___closed__2); -l_Lean_Parser_Command_def___closed__3 = _init_l_Lean_Parser_Command_def___closed__3(); -lean_mark_persistent(l_Lean_Parser_Command_def___closed__3); -l_Lean_Parser_Command_def___closed__4 = _init_l_Lean_Parser_Command_def___closed__4(); -lean_mark_persistent(l_Lean_Parser_Command_def___closed__4); -l_Lean_Parser_Command_def___closed__5 = _init_l_Lean_Parser_Command_def___closed__5(); -lean_mark_persistent(l_Lean_Parser_Command_def___closed__5); -l_Lean_Parser_Command_def___closed__6 = _init_l_Lean_Parser_Command_def___closed__6(); -lean_mark_persistent(l_Lean_Parser_Command_def___closed__6); -l_Lean_Parser_Command_def = _init_l_Lean_Parser_Command_def(); -lean_mark_persistent(l_Lean_Parser_Command_def); l_Lean_Parser_Command_theorem___closed__1 = _init_l_Lean_Parser_Command_theorem___closed__1(); lean_mark_persistent(l_Lean_Parser_Command_theorem___closed__1); l_Lean_Parser_Command_theorem___closed__2 = _init_l_Lean_Parser_Command_theorem___closed__2(); @@ -44374,8 +44083,6 @@ l_Lean_Parser_Command_declaration___closed__16 = _init_l_Lean_Parser_Command_dec lean_mark_persistent(l_Lean_Parser_Command_declaration___closed__16); l_Lean_Parser_Command_declaration___closed__17 = _init_l_Lean_Parser_Command_declaration___closed__17(); lean_mark_persistent(l_Lean_Parser_Command_declaration___closed__17); -l_Lean_Parser_Command_declaration___closed__18 = _init_l_Lean_Parser_Command_declaration___closed__18(); -lean_mark_persistent(l_Lean_Parser_Command_declaration___closed__18); l_Lean_Parser_Command_declaration = _init_l_Lean_Parser_Command_declaration(); lean_mark_persistent(l_Lean_Parser_Command_declaration); if (builtin) {res = l___regBuiltin_Lean_Parser_Command_declaration(lean_io_mk_world()); @@ -44734,17 +44441,6 @@ lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_definition_formatter___c if (builtin) {res = l___regBuiltin_Lean_Parser_Command_definition_formatter(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}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(); -lean_mark_persistent(l_Lean_Parser_Command_def_formatter___closed__2); -l___regBuiltin_Lean_Parser_Command_def_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_def_formatter___closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_def_formatter___closed__1); -l___regBuiltin_Lean_Parser_Command_def_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_def_formatter___closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_def_formatter___closed__2); -if (builtin) {res = l___regBuiltin_Lean_Parser_Command_def_formatter(lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); }l_Lean_Parser_Command_declSig_formatter___closed__1 = _init_l_Lean_Parser_Command_declSig_formatter___closed__1(); lean_mark_persistent(l_Lean_Parser_Command_declSig_formatter___closed__1); l_Lean_Parser_Command_declSig_formatter___closed__2 = _init_l_Lean_Parser_Command_declSig_formatter___closed__2(); @@ -45330,8 +45026,6 @@ l_Lean_Parser_Command_declaration_formatter___closed__12 = _init_l_Lean_Parser_C lean_mark_persistent(l_Lean_Parser_Command_declaration_formatter___closed__12); l_Lean_Parser_Command_declaration_formatter___closed__13 = _init_l_Lean_Parser_Command_declaration_formatter___closed__13(); lean_mark_persistent(l_Lean_Parser_Command_declaration_formatter___closed__13); -l_Lean_Parser_Command_declaration_formatter___closed__14 = _init_l_Lean_Parser_Command_declaration_formatter___closed__14(); -lean_mark_persistent(l_Lean_Parser_Command_declaration_formatter___closed__14); l___regBuiltin_Lean_Parser_Command_declaration_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_declaration_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_declaration_formatter___closed__1); l___regBuiltin_Lean_Parser_Command_declaration_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_declaration_formatter___closed__2(); @@ -45675,17 +45369,6 @@ lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_definition_parenthesizer if (builtin) {res = l___regBuiltin_Lean_Parser_Command_definition_parenthesizer(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}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(); -lean_mark_persistent(l_Lean_Parser_Command_def_parenthesizer___closed__2); -l___regBuiltin_Lean_Parser_Command_def_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_def_parenthesizer___closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_def_parenthesizer___closed__1); -l___regBuiltin_Lean_Parser_Command_def_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_def_parenthesizer___closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_def_parenthesizer___closed__2); -if (builtin) {res = l___regBuiltin_Lean_Parser_Command_def_parenthesizer(lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); }l_Lean_Parser_Command_declSig_parenthesizer___closed__1 = _init_l_Lean_Parser_Command_declSig_parenthesizer___closed__1(); lean_mark_persistent(l_Lean_Parser_Command_declSig_parenthesizer___closed__1); l_Lean_Parser_Command_declSig_parenthesizer___closed__2 = _init_l_Lean_Parser_Command_declSig_parenthesizer___closed__2(); @@ -46271,8 +45954,6 @@ l_Lean_Parser_Command_declaration_parenthesizer___closed__12 = _init_l_Lean_Pars lean_mark_persistent(l_Lean_Parser_Command_declaration_parenthesizer___closed__12); l_Lean_Parser_Command_declaration_parenthesizer___closed__13 = _init_l_Lean_Parser_Command_declaration_parenthesizer___closed__13(); lean_mark_persistent(l_Lean_Parser_Command_declaration_parenthesizer___closed__13); -l_Lean_Parser_Command_declaration_parenthesizer___closed__14 = _init_l_Lean_Parser_Command_declaration_parenthesizer___closed__14(); -lean_mark_persistent(l_Lean_Parser_Command_declaration_parenthesizer___closed__14); l___regBuiltin_Lean_Parser_Command_declaration_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_declaration_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_declaration_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Command_declaration_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_declaration_parenthesizer___closed__2(); @@ -49062,122 +48743,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); -if (builtin) {res = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2829_(lean_io_mk_world()); +if (builtin) {res = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2791_(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_2851____closed__1 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__1(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__1); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__2 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__2(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__2); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__3 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__3(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__3); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__4 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__4(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__4); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__5 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__5(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__5); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__6 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__6(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__6); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__7 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__7(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__7); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__8 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__8(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__8); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__9 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__9(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__9); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__10 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__10(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__10); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__11 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__11(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__11); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__12 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__12(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__12); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__13 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__13(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__13); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__14 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__14(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__14); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__15 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__15(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__15); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__16 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__16(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__16); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__17 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__17(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__17); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__18 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__18(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__18); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__19 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__19(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__19); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__20 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__20(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__20); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__21 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__21(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__21); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__22 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__22(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__22); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__23 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__23(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__23); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__24 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__24(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__24); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__25 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__25(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__25); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__26 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__26(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__26); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__27 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__27(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__27); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__28 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__28(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__28); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__29 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__29(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__29); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__30 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__30(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__30); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__31 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__31(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__31); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__32 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__32(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__32); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__33 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__33(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__33); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__34 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__34(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__34); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__35 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__35(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__35); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__36 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__36(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__36); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__37 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__37(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__37); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__38 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__38(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__38); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__39 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__39(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__39); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__40 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__40(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__40); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__41 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__41(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__41); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__42 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__42(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__42); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__43 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__43(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__43); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__44 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__44(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__44); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__45 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__45(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__45); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__46 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__46(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__46); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__47 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__47(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__47); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__48 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__48(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__48); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__49 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__49(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__49); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__50 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__50(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__50); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__51 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__51(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__51); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__52 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__52(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__52); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__53 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__53(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__53); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__54 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__54(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851____closed__54); -if (builtin) {res = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2851_(lean_io_mk_world()); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__1 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__1(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__1); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__2 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__2(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__2); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__3 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__3(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__3); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__4 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__4(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__4); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__5 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__5(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__5); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__6 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__6(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__6); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__7 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__7(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__7); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__8 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__8(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__8); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__9 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__9(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__9); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__10 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__10(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__10); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__11 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__11(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__11); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__12 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__12(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__12); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__13 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__13(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__13); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__14 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__14(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__14); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__15 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__15(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__15); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__16 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__16(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__16); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__17 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__17(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__17); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__18 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__18(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__18); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__19 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__19(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__19); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__20 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__20(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__20); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__21 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__21(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__21); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__22 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__22(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__22); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__23 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__23(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__23); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__24 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__24(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__24); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__25 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__25(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__25); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__26 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__26(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__26); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__27 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__27(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__27); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__28 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__28(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__28); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__29 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__29(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__29); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__30 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__30(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__30); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__31 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__31(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__31); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__32 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__32(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__32); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__33 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__33(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__33); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__34 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__34(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__34); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__35 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__35(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__35); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__36 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__36(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__36); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__37 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__37(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__37); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__38 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__38(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__38); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__39 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__39(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__39); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__40 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__40(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__40); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__41 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__41(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__41); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__42 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__42(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__42); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__43 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__43(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__43); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__44 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__44(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__44); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__45 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__45(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__45); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__46 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__46(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__46); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__47 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__47(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__47); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__48 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__48(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__48); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__49 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__49(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__49); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__50 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__50(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__50); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__51 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__51(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__51); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__52 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__52(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__52); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__53 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__53(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__53); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__54 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__54(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813____closed__54); +if (builtin) {res = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2813_(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(); From 995726f75fd7eed223c2189a54f6df3293185327 Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Wed, 13 Mar 2024 20:57:53 -0700 Subject: [PATCH 32/32] chore: fix tests --- tests/lean/StxQuot.lean.expected.out | 8 ++++---- tests/plugin/SnakeLinter.lean | 4 ++-- 2 files changed, 6 insertions(+), 6 deletions(-) diff --git a/tests/lean/StxQuot.lean.expected.out b/tests/lean/StxQuot.lean.expected.out index a038be803a60..2411f8e1a804 100644 --- a/tests/lean/StxQuot.lean.expected.out +++ b/tests/lean/StxQuot.lean.expected.out @@ -8,8 +8,8 @@ StxQuot.lean:8:12-8:13: error: unexpected token ')'; expected identifier or term "(«term_+_» (num \"1\") \"+\" (num \"1\"))" StxQuot.lean:19:15-19:16: error: unexpected token ']'; expected term "(Term.fun \"fun\" (Term.basicFun [`a._@.UnhygienicMain._hyg.1] [] \"=>\" `a._@.UnhygienicMain._hyg.1))" -"(Command.declaration\n (Command.declModifiers [] [] [] [] [] [])\n (Command.def\n \"def\"\n (Command.declId `foo._@.UnhygienicMain._hyg.1 [])\n (Command.optDeclSig [] [])\n (Command.declValSimple \":=\" (num \"1\") (Termination.suffix [] []) [])\n []))" -"[(Command.declaration\n (Command.declModifiers [] [] [] [] [] [])\n (Command.def\n \"def\"\n (Command.declId `foo._@.UnhygienicMain._hyg.1 [])\n (Command.optDeclSig [] [])\n (Command.declValSimple \":=\" (num \"1\") (Termination.suffix [] []) [])\n []))\n (Command.declaration\n (Command.declModifiers [] [] [] [] [] [])\n (Command.def\n \"def\"\n (Command.declId `bar._@.UnhygienicMain._hyg.1 [])\n (Command.optDeclSig [] [])\n (Command.declValSimple \":=\" (num \"2\") (Termination.suffix [] []) [])\n []))]" +"(Command.declaration\n (Command.declModifiers [] [] [] [] [] [])\n (Command.definition\n \"def\"\n (Command.declId `foo._@.UnhygienicMain._hyg.1 [])\n (Command.optDeclSig [] [])\n (Command.declValSimple \":=\" (num \"1\") (Termination.suffix [] []) [])\n []))" +"[(Command.declaration\n (Command.declModifiers [] [] [] [] [] [])\n (Command.definition\n \"def\"\n (Command.declId `foo._@.UnhygienicMain._hyg.1 [])\n (Command.optDeclSig [] [])\n (Command.declValSimple \":=\" (num \"1\") (Termination.suffix [] []) [])\n []))\n (Command.declaration\n (Command.declModifiers [] [] [] [] [] [])\n (Command.definition\n \"def\"\n (Command.declId `bar._@.UnhygienicMain._hyg.1 [])\n (Command.optDeclSig [] [])\n (Command.declValSimple \":=\" (num \"2\") (Termination.suffix [] []) [])\n []))]" "`Nat.one._@.UnhygienicMain._hyg.1" "`Nat.one._@.UnhygienicMain._hyg.1" "(Term.app `f._@.UnhygienicMain._hyg.1 [`Nat.one._@.UnhygienicMain._hyg.1 `Nat.one._@.UnhygienicMain._hyg.1])" @@ -18,8 +18,8 @@ StxQuot.lean:19:15-19:16: error: unexpected token ']'; expected term "(Term.proj `Nat.one._@.UnhygienicMain._hyg.1 \".\" `b._@.UnhygienicMain._hyg.1)" "(«term_+_» (num \"2\") \"+\" (num \"1\"))" "(«term_+_» («term_+_» (num \"1\") \"+\" (num \"2\")) \"+\" (num \"1\"))" -"(Command.declaration\n (Command.declModifiers [] [] [] [] [] [])\n (Command.def\n \"def\"\n (Command.declId `foo._@.UnhygienicMain._hyg.1 [])\n (Command.optDeclSig [] [])\n (Command.declValSimple \":=\" (num \"1\") (Termination.suffix [] []) [])\n []))" -"[(Command.declaration\n (Command.declModifiers [] [] [] [] [] [])\n (Command.def\n \"def\"\n (Command.declId `bar._@.UnhygienicMain._hyg.1 [])\n (Command.optDeclSig [] [])\n (Command.declValSimple \":=\" (num \"2\") (Termination.suffix [] []) [])\n []))\n (Command.declaration\n (Command.declModifiers [] [] [] [] [] [])\n (Command.def\n \"def\"\n (Command.declId `foo._@.UnhygienicMain._hyg.1 [])\n (Command.optDeclSig [] [])\n (Command.declValSimple \":=\" (num \"1\") (Termination.suffix [] []) [])\n []))]" +"(Command.declaration\n (Command.declModifiers [] [] [] [] [] [])\n (Command.definition\n \"def\"\n (Command.declId `foo._@.UnhygienicMain._hyg.1 [])\n (Command.optDeclSig [] [])\n (Command.declValSimple \":=\" (num \"1\") (Termination.suffix [] []) [])\n []))" +"[(Command.declaration\n (Command.declModifiers [] [] [] [] [] [])\n (Command.definition\n \"def\"\n (Command.declId `bar._@.UnhygienicMain._hyg.1 [])\n (Command.optDeclSig [] [])\n (Command.declValSimple \":=\" (num \"2\") (Termination.suffix [] []) [])\n []))\n (Command.declaration\n (Command.declModifiers [] [] [] [] [] [])\n (Command.definition\n \"def\"\n (Command.declId `foo._@.UnhygienicMain._hyg.1 [])\n (Command.optDeclSig [] [])\n (Command.declValSimple \":=\" (num \"1\") (Termination.suffix [] []) [])\n []))]" "0" 0 1 diff --git a/tests/plugin/SnakeLinter.lean b/tests/plugin/SnakeLinter.lean index 93a0a77d6962..093e15fd8f3b 100644 --- a/tests/plugin/SnakeLinter.lean +++ b/tests/plugin/SnakeLinter.lean @@ -5,9 +5,9 @@ def oh_no : Nat := 0 def snakeLinter : Linter where run stx := do - if stx.getKind == `Lean.Parser.Command.declaration then + if stx.getKind == ``Lean.Parser.Command.declaration then let decl := stx[1] - if decl.getKind == `Lean.Parser.Command.def then + if decl.getKind == ``Lean.Parser.Command.definition then let declId := decl[1] withRef declId do let declName := declId[0].getId